src/HOL/List.thy
author bulwahn
Thu, 08 Dec 2011 13:53:27 +0100
changeset 45789 36ea69266e61
parent 45714 ad4242285560
child 45794 16e8e4d33c42
permissions -rw-r--r--
removing special code generator setup for hd and last function because this causes problems with quickcheck narrowing as the Haskell Prelude functions throw errors that cannot be caught instead of PatternFail exceptions
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 13366
diff changeset
     1
(*  Title:      HOL/List.thy
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 13366
diff changeset
     2
    Author:     Tobias Nipkow
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     3
*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     4
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
     5
header {* The datatype of finite lists *}
13122
wenzelm
parents: 13114
diff changeset
     6
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15113
diff changeset
     7
theory List
44013
5cfc1c36ae97 moved recdef package to HOL/Library/Old_Recdef.thy
krauss
parents: 43594
diff changeset
     8
imports Plain Presburger Code_Numeral Quotient ATP
41463
edbf0a86fb1c adding simproc to rewrite list comprehensions to set comprehensions; adopting proofs
bulwahn
parents: 41372
diff changeset
     9
uses
edbf0a86fb1c adding simproc to rewrite list comprehensions to set comprehensions; adopting proofs
bulwahn
parents: 41372
diff changeset
    10
  ("Tools/list_code.ML")
edbf0a86fb1c adding simproc to rewrite list comprehensions to set comprehensions; adopting proofs
bulwahn
parents: 41372
diff changeset
    11
  ("Tools/list_to_set_comprehension.ML")
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15113
diff changeset
    12
begin
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    13
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
    14
datatype 'a list =
13366
114b7c14084a moved stuff from Main.thy;
wenzelm
parents: 13187
diff changeset
    15
    Nil    ("[]")
114b7c14084a moved stuff from Main.thy;
wenzelm
parents: 13187
diff changeset
    16
  | Cons 'a  "'a list"    (infixr "#" 65)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    17
34941
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    18
syntax
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    19
  -- {* list Enumeration *}
35115
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
    20
  "_list" :: "args => 'a list"    ("[(_)]")
34941
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    21
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    22
translations
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    23
  "[x, xs]" == "x#[xs]"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    24
  "[x]" == "x#[]"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    25
35115
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
    26
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
    27
subsection {* Basic list processing functions *}
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
    28
34941
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    29
primrec
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    30
  hd :: "'a list \<Rightarrow> 'a" where
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    31
  "hd (x # xs) = x"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    32
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    33
primrec
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    34
  tl :: "'a list \<Rightarrow> 'a list" where
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    35
    "tl [] = []"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    36
  | "tl (x # xs) = xs"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    37
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    38
primrec
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    39
  last :: "'a list \<Rightarrow> 'a" where
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    40
  "last (x # xs) = (if xs = [] then x else last xs)"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    41
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    42
primrec
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    43
  butlast :: "'a list \<Rightarrow> 'a list" where
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    44
    "butlast []= []"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    45
  | "butlast (x # xs) = (if xs = [] then [] else x # butlast xs)"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    46
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    47
primrec
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    48
  set :: "'a list \<Rightarrow> 'a set" where
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    49
    "set [] = {}"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    50
  | "set (x # xs) = insert x (set xs)"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    51
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    52
primrec
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    53
  map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b list" where
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    54
    "map f [] = []"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    55
  | "map f (x # xs) = f x # map f xs"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    56
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    57
primrec
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    58
  append :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infixr "@" 65) where
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    59
    append_Nil:"[] @ ys = ys"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    60
  | append_Cons: "(x#xs) @ ys = x # xs @ ys"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    61
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    62
primrec
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    63
  rev :: "'a list \<Rightarrow> 'a list" where
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    64
    "rev [] = []"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    65
  | "rev (x # xs) = rev xs @ [x]"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    66
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    67
primrec
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    68
  filter:: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list" where
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    69
    "filter P [] = []"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    70
  | "filter P (x # xs) = (if P x then x # filter P xs else filter P xs)"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    71
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    72
syntax
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    73
  -- {* Special syntax for filter *}
35115
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
    74
  "_filter" :: "[pttrn, 'a list, bool] => 'a list"    ("(1[_<-_./ _])")
34941
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    75
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    76
translations
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    77
  "[x<-xs . P]"== "CONST filter (%x. P) xs"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    78
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    79
syntax (xsymbols)
35115
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
    80
  "_filter" :: "[pttrn, 'a list, bool] => 'a list"("(1[_\<leftarrow>_ ./ _])")
34941
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    81
syntax (HTML output)
35115
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
    82
  "_filter" :: "[pttrn, 'a list, bool] => 'a list"("(1[_\<leftarrow>_ ./ _])")
34941
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    83
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    84
primrec
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    85
  foldl :: "('b \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a list \<Rightarrow> 'b" where
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    86
    foldl_Nil: "foldl f a [] = a"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    87
  | foldl_Cons: "foldl f a (x # xs) = foldl f (f a x) xs"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    88
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    89
primrec
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    90
  foldr :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b \<Rightarrow> 'b" where
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    91
    "foldr f [] a = a"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    92
  | "foldr f (x # xs) a = f x (foldr f xs a)"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    93
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    94
primrec
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    95
  concat:: "'a list list \<Rightarrow> 'a list" where
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    96
    "concat [] = []"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    97
  | "concat (x # xs) = x @ concat xs"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
    98
39774
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
    99
definition (in monoid_add)
34941
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   100
  listsum :: "'a list \<Rightarrow> 'a" where
39774
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
   101
  "listsum xs = foldr plus xs 0"
34941
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   102
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   103
primrec
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   104
  drop:: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list" where
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   105
    drop_Nil: "drop n [] = []"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   106
  | drop_Cons: "drop n (x # xs) = (case n of 0 \<Rightarrow> x # xs | Suc m \<Rightarrow> drop m xs)"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   107
  -- {*Warning: simpset does not contain this definition, but separate
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   108
       theorems for @{text "n = 0"} and @{text "n = Suc k"} *}
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   109
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   110
primrec
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   111
  take:: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list" where
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   112
    take_Nil:"take n [] = []"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   113
  | take_Cons: "take n (x # xs) = (case n of 0 \<Rightarrow> [] | Suc m \<Rightarrow> x # take m xs)"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   114
  -- {*Warning: simpset does not contain this definition, but separate
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   115
       theorems for @{text "n = 0"} and @{text "n = Suc k"} *}
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   116
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   117
primrec
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   118
  nth :: "'a list => nat => 'a" (infixl "!" 100) where
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   119
  nth_Cons: "(x # xs) ! n = (case n of 0 \<Rightarrow> x | Suc k \<Rightarrow> xs ! k)"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   120
  -- {*Warning: simpset does not contain this definition, but separate
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   121
       theorems for @{text "n = 0"} and @{text "n = Suc k"} *}
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   122
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   123
primrec
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   124
  list_update :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a list" where
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   125
    "list_update [] i v = []"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   126
  | "list_update (x # xs) i v = (case i of 0 \<Rightarrow> v # xs | Suc j \<Rightarrow> x # list_update xs j v)"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   127
41229
d797baa3d57c replaced command 'nonterminals' by slightly modernized version 'nonterminal';
wenzelm
parents: 41075
diff changeset
   128
nonterminal lupdbinds and lupdbind
5077
71043526295f * HOL/List: new function list_update written xs[i:=v] that updates the i-th
nipkow
parents: 4643
diff changeset
   129
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   130
syntax
13366
114b7c14084a moved stuff from Main.thy;
wenzelm
parents: 13187
diff changeset
   131
  "_lupdbind":: "['a, 'a] => lupdbind"    ("(2_ :=/ _)")
114b7c14084a moved stuff from Main.thy;
wenzelm
parents: 13187
diff changeset
   132
  "" :: "lupdbind => lupdbinds"    ("_")
114b7c14084a moved stuff from Main.thy;
wenzelm
parents: 13187
diff changeset
   133
  "_lupdbinds" :: "[lupdbind, lupdbinds] => lupdbinds"    ("_,/ _")
114b7c14084a moved stuff from Main.thy;
wenzelm
parents: 13187
diff changeset
   134
  "_LUpdate" :: "['a, lupdbinds] => 'a"    ("_/[(_)]" [900,0] 900)
5077
71043526295f * HOL/List: new function list_update written xs[i:=v] that updates the i-th
nipkow
parents: 4643
diff changeset
   135
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   136
translations
35115
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
   137
  "_LUpdate xs (_lupdbinds b bs)" == "_LUpdate (_LUpdate xs b) bs"
34941
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   138
  "xs[i:=x]" == "CONST list_update xs i x"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   139
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   140
primrec
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   141
  takeWhile :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list" where
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   142
    "takeWhile P [] = []"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   143
  | "takeWhile P (x # xs) = (if P x then x # takeWhile P xs else [])"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   144
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   145
primrec
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   146
  dropWhile :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list" where
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   147
    "dropWhile P [] = []"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   148
  | "dropWhile P (x # xs) = (if P x then dropWhile P xs else x # xs)"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   149
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   150
primrec
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   151
  zip :: "'a list \<Rightarrow> 'b list \<Rightarrow> ('a \<times> 'b) list" where
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   152
    "zip xs [] = []"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   153
  | zip_Cons: "zip xs (y # ys) = (case xs of [] => [] | z # zs => (z, y) # zip zs ys)"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   154
  -- {*Warning: simpset does not contain this definition, but separate
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   155
       theorems for @{text "xs = []"} and @{text "xs = z # zs"} *}
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   156
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   157
primrec 
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   158
  upt :: "nat \<Rightarrow> nat \<Rightarrow> nat list" ("(1[_..</_'])") where
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   159
    upt_0: "[i..<0] = []"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   160
  | upt_Suc: "[i..<(Suc j)] = (if i <= j then [i..<j] @ [j] else [])"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   161
34978
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
   162
definition
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
   163
  insert :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
   164
  "insert x xs = (if x \<in> set xs then xs else x # xs)"
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
   165
36176
3fe7e97ccca8 replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
wenzelm
parents: 36154
diff changeset
   166
hide_const (open) insert
3fe7e97ccca8 replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
wenzelm
parents: 36154
diff changeset
   167
hide_fact (open) insert_def
34978
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
   168
34941
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   169
primrec
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   170
  remove1 :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   171
    "remove1 x [] = []"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   172
  | "remove1 x (y # xs) = (if x = y then xs else y # remove1 x xs)"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   173
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   174
primrec
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   175
  removeAll :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   176
    "removeAll x [] = []"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   177
  | "removeAll x (y # xs) = (if x = y then removeAll x xs else y # removeAll x xs)"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   178
40122
1d8ad2ff3e01 dropped (almost) redundant distinct.induct rule; distinct_simps again named distinct.simps
haftmann
parents: 40077
diff changeset
   179
primrec
39915
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
   180
  distinct :: "'a list \<Rightarrow> bool" where
40122
1d8ad2ff3e01 dropped (almost) redundant distinct.induct rule; distinct_simps again named distinct.simps
haftmann
parents: 40077
diff changeset
   181
    "distinct [] \<longleftrightarrow> True"
1d8ad2ff3e01 dropped (almost) redundant distinct.induct rule; distinct_simps again named distinct.simps
haftmann
parents: 40077
diff changeset
   182
  | "distinct (x # xs) \<longleftrightarrow> x \<notin> set xs \<and> distinct xs"
39915
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
   183
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
   184
primrec
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
   185
  remdups :: "'a list \<Rightarrow> 'a list" where
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
   186
    "remdups [] = []"
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
   187
  | "remdups (x # xs) = (if x \<in> set xs then remdups xs else x # remdups xs)"
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
   188
34941
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   189
primrec
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   190
  replicate :: "nat \<Rightarrow> 'a \<Rightarrow> 'a list" where
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   191
    replicate_0: "replicate 0 x = []"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   192
  | replicate_Suc: "replicate (Suc n) x = x # replicate n x"
3342
ec3b55fcb165 New operator "lists" for formalizing sets of lists
paulson
parents: 3320
diff changeset
   193
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   194
text {*
14589
feae7b5fd425 tuned document;
wenzelm
parents: 14565
diff changeset
   195
  Function @{text size} is overloaded for all datatypes. Users may
13366
114b7c14084a moved stuff from Main.thy;
wenzelm
parents: 13187
diff changeset
   196
  refer to the list version as @{text length}. *}
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   197
19363
667b5ea637dd refined 'abbreviation';
wenzelm
parents: 19302
diff changeset
   198
abbreviation
34941
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   199
  length :: "'a list \<Rightarrow> nat" where
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
   200
  "length \<equiv> size"
15307
10dd989282fd indentation
paulson
parents: 15305
diff changeset
   201
21061
580dfc999ef6 added normal post setup; cleaned up "execution" constants
haftmann
parents: 21046
diff changeset
   202
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21211
diff changeset
   203
  rotate1 :: "'a list \<Rightarrow> 'a list" where
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21211
diff changeset
   204
  "rotate1 xs = (case xs of [] \<Rightarrow> [] | x#xs \<Rightarrow> xs @ [x])"
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21211
diff changeset
   205
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21211
diff changeset
   206
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21211
diff changeset
   207
  rotate :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list" where
30971
7fbebf75b3ef funpow and relpow with shared "^^" syntax
haftmann
parents: 30952
diff changeset
   208
  "rotate n = rotate1 ^^ n"
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21211
diff changeset
   209
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21211
diff changeset
   210
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21211
diff changeset
   211
  list_all2 :: "('a => 'b => bool) => 'a list => 'b list => bool" where
37767
a2b7a20d6ea3 dropped superfluous [code del]s
haftmann
parents: 37605
diff changeset
   212
  "list_all2 P xs ys =
21061
580dfc999ef6 added normal post setup; cleaned up "execution" constants
haftmann
parents: 21046
diff changeset
   213
    (length xs = length ys \<and> (\<forall>(x, y) \<in> set (zip xs ys). P x y))"
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21211
diff changeset
   214
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21211
diff changeset
   215
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21211
diff changeset
   216
  sublist :: "'a list => nat set => 'a list" where
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21211
diff changeset
   217
  "sublist xs A = map fst (filter (\<lambda>p. snd p \<in> A) (zip xs [0..<size xs]))"
17086
0eb0c9259dd7 added quite a few functions for code generation
nipkow
parents: 16998
diff changeset
   218
40593
1e57b18d27b1 code eqn for slice was missing; redefined splice with fun
nipkow
parents: 40365
diff changeset
   219
fun splice :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
1e57b18d27b1 code eqn for slice was missing; redefined splice with fun
nipkow
parents: 40365
diff changeset
   220
"splice [] ys = ys" |
1e57b18d27b1 code eqn for slice was missing; redefined splice with fun
nipkow
parents: 40365
diff changeset
   221
"splice xs [] = xs" |
1e57b18d27b1 code eqn for slice was missing; redefined splice with fun
nipkow
parents: 40365
diff changeset
   222
"splice (x#xs) (y#ys) = x # y # splice xs ys"
21061
580dfc999ef6 added normal post setup; cleaned up "execution" constants
haftmann
parents: 21046
diff changeset
   223
26771
1d67ab20f358 Added documentation
nipkow
parents: 26749
diff changeset
   224
text{*
1d67ab20f358 Added documentation
nipkow
parents: 26749
diff changeset
   225
\begin{figure}[htbp]
1d67ab20f358 Added documentation
nipkow
parents: 26749
diff changeset
   226
\fbox{
1d67ab20f358 Added documentation
nipkow
parents: 26749
diff changeset
   227
\begin{tabular}{l}
27381
19ae7064f00f @{lemma}: 'by' keyword;
wenzelm
parents: 27368
diff changeset
   228
@{lemma "[a,b]@[c,d] = [a,b,c,d]" by simp}\\
19ae7064f00f @{lemma}: 'by' keyword;
wenzelm
parents: 27368
diff changeset
   229
@{lemma "length [a,b,c] = 3" by simp}\\
19ae7064f00f @{lemma}: 'by' keyword;
wenzelm
parents: 27368
diff changeset
   230
@{lemma "set [a,b,c] = {a,b,c}" by simp}\\
19ae7064f00f @{lemma}: 'by' keyword;
wenzelm
parents: 27368
diff changeset
   231
@{lemma "map f [a,b,c] = [f a, f b, f c]" by simp}\\
19ae7064f00f @{lemma}: 'by' keyword;
wenzelm
parents: 27368
diff changeset
   232
@{lemma "rev [a,b,c] = [c,b,a]" by simp}\\
19ae7064f00f @{lemma}: 'by' keyword;
wenzelm
parents: 27368
diff changeset
   233
@{lemma "hd [a,b,c,d] = a" by simp}\\
19ae7064f00f @{lemma}: 'by' keyword;
wenzelm
parents: 27368
diff changeset
   234
@{lemma "tl [a,b,c,d] = [b,c,d]" by simp}\\
19ae7064f00f @{lemma}: 'by' keyword;
wenzelm
parents: 27368
diff changeset
   235
@{lemma "last [a,b,c,d] = d" by simp}\\
19ae7064f00f @{lemma}: 'by' keyword;
wenzelm
parents: 27368
diff changeset
   236
@{lemma "butlast [a,b,c,d] = [a,b,c]" by simp}\\
19ae7064f00f @{lemma}: 'by' keyword;
wenzelm
parents: 27368
diff changeset
   237
@{lemma[source] "filter (\<lambda>n::nat. n<2) [0,2,1] = [0,1]" by simp}\\
19ae7064f00f @{lemma}: 'by' keyword;
wenzelm
parents: 27368
diff changeset
   238
@{lemma "concat [[a,b],[c,d,e],[],[f]] = [a,b,c,d,e,f]" by simp}\\
19ae7064f00f @{lemma}: 'by' keyword;
wenzelm
parents: 27368
diff changeset
   239
@{lemma "foldl f x [a,b,c] = f (f (f x a) b) c" by simp}\\
19ae7064f00f @{lemma}: 'by' keyword;
wenzelm
parents: 27368
diff changeset
   240
@{lemma "foldr f [a,b,c] x = f a (f b (f c x))" by simp}\\
19ae7064f00f @{lemma}: 'by' keyword;
wenzelm
parents: 27368
diff changeset
   241
@{lemma "zip [a,b,c] [x,y,z] = [(a,x),(b,y),(c,z)]" by simp}\\
19ae7064f00f @{lemma}: 'by' keyword;
wenzelm
parents: 27368
diff changeset
   242
@{lemma "zip [a,b] [x,y,z] = [(a,x),(b,y)]" by simp}\\
19ae7064f00f @{lemma}: 'by' keyword;
wenzelm
parents: 27368
diff changeset
   243
@{lemma "splice [a,b,c] [x,y,z] = [a,x,b,y,c,z]" by simp}\\
19ae7064f00f @{lemma}: 'by' keyword;
wenzelm
parents: 27368
diff changeset
   244
@{lemma "splice [a,b,c,d] [x,y] = [a,x,b,y,c,d]" by simp}\\
19ae7064f00f @{lemma}: 'by' keyword;
wenzelm
parents: 27368
diff changeset
   245
@{lemma "take 2 [a,b,c,d] = [a,b]" by simp}\\
19ae7064f00f @{lemma}: 'by' keyword;
wenzelm
parents: 27368
diff changeset
   246
@{lemma "take 6 [a,b,c,d] = [a,b,c,d]" by simp}\\
19ae7064f00f @{lemma}: 'by' keyword;
wenzelm
parents: 27368
diff changeset
   247
@{lemma "drop 2 [a,b,c,d] = [c,d]" by simp}\\
19ae7064f00f @{lemma}: 'by' keyword;
wenzelm
parents: 27368
diff changeset
   248
@{lemma "drop 6 [a,b,c,d] = []" by simp}\\
19ae7064f00f @{lemma}: 'by' keyword;
wenzelm
parents: 27368
diff changeset
   249
@{lemma "takeWhile (%n::nat. n<3) [1,2,3,0] = [1,2]" by simp}\\
19ae7064f00f @{lemma}: 'by' keyword;
wenzelm
parents: 27368
diff changeset
   250
@{lemma "dropWhile (%n::nat. n<3) [1,2,3,0] = [3,0]" by simp}\\
19ae7064f00f @{lemma}: 'by' keyword;
wenzelm
parents: 27368
diff changeset
   251
@{lemma "distinct [2,0,1::nat]" by simp}\\
19ae7064f00f @{lemma}: 'by' keyword;
wenzelm
parents: 27368
diff changeset
   252
@{lemma "remdups [2,0,2,1::nat,2] = [0,1,2]" by simp}\\
34978
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
   253
@{lemma "List.insert 2 [0::nat,1,2] = [0,1,2]" by (simp add: List.insert_def)}\\
35295
397295fa8387 lemma distinct_insert
haftmann
parents: 35248
diff changeset
   254
@{lemma "List.insert 3 [0::nat,1,2] = [3,0,1,2]" by (simp add: List.insert_def)}\\
27381
19ae7064f00f @{lemma}: 'by' keyword;
wenzelm
parents: 27368
diff changeset
   255
@{lemma "remove1 2 [2,0,2,1::nat,2] = [0,2,1,2]" by simp}\\
27693
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
   256
@{lemma "removeAll 2 [2,0,2,1::nat,2] = [0,1]" by simp}\\
27381
19ae7064f00f @{lemma}: 'by' keyword;
wenzelm
parents: 27368
diff changeset
   257
@{lemma "nth [a,b,c,d] 2 = c" by simp}\\
19ae7064f00f @{lemma}: 'by' keyword;
wenzelm
parents: 27368
diff changeset
   258
@{lemma "[a,b,c,d][2 := x] = [a,b,x,d]" by simp}\\
19ae7064f00f @{lemma}: 'by' keyword;
wenzelm
parents: 27368
diff changeset
   259
@{lemma "sublist [a,b,c,d,e] {0,2,3} = [a,c,d]" by (simp add:sublist_def)}\\
19ae7064f00f @{lemma}: 'by' keyword;
wenzelm
parents: 27368
diff changeset
   260
@{lemma "rotate1 [a,b,c,d] = [b,c,d,a]" by (simp add:rotate1_def)}\\
40077
c8a9eaaa2f59 nat_number -> eval_nat_numeral
nipkow
parents: 39963
diff changeset
   261
@{lemma "rotate 3 [a,b,c,d] = [d,a,b,c]" by (simp add:rotate1_def rotate_def eval_nat_numeral)}\\
c8a9eaaa2f59 nat_number -> eval_nat_numeral
nipkow
parents: 39963
diff changeset
   262
@{lemma "replicate 4 a = [a,a,a,a]" by (simp add:eval_nat_numeral)}\\
c8a9eaaa2f59 nat_number -> eval_nat_numeral
nipkow
parents: 39963
diff changeset
   263
@{lemma "[2..<5] = [2,3,4]" by (simp add:eval_nat_numeral)}\\
39774
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
   264
@{lemma "listsum [1,2,3::nat] = 6" by (simp add: listsum_def)}
26771
1d67ab20f358 Added documentation
nipkow
parents: 26749
diff changeset
   265
\end{tabular}}
1d67ab20f358 Added documentation
nipkow
parents: 26749
diff changeset
   266
\caption{Characteristic examples}
1d67ab20f358 Added documentation
nipkow
parents: 26749
diff changeset
   267
\label{fig:Characteristic}
1d67ab20f358 Added documentation
nipkow
parents: 26749
diff changeset
   268
\end{figure}
29927
ae8f42c245b2 Added nitpick attribute, and fixed typo.
blanchet
parents: 29856
diff changeset
   269
Figure~\ref{fig:Characteristic} shows characteristic examples
26771
1d67ab20f358 Added documentation
nipkow
parents: 26749
diff changeset
   270
that should give an intuitive understanding of the above functions.
1d67ab20f358 Added documentation
nipkow
parents: 26749
diff changeset
   271
*}
1d67ab20f358 Added documentation
nipkow
parents: 26749
diff changeset
   272
24616
fac3dd4ade83 sorting
nipkow
parents: 24566
diff changeset
   273
text{* The following simple sort functions are intended for proofs,
fac3dd4ade83 sorting
nipkow
parents: 24566
diff changeset
   274
not for efficient implementations. *}
fac3dd4ade83 sorting
nipkow
parents: 24566
diff changeset
   275
25221
5ded95dda5df append/member: more light-weight way to declare authentic syntax;
wenzelm
parents: 25215
diff changeset
   276
context linorder
5ded95dda5df append/member: more light-weight way to declare authentic syntax;
wenzelm
parents: 25215
diff changeset
   277
begin
5ded95dda5df append/member: more light-weight way to declare authentic syntax;
wenzelm
parents: 25215
diff changeset
   278
39915
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
   279
inductive sorted :: "'a list \<Rightarrow> bool" where
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
   280
  Nil [iff]: "sorted []"
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
   281
| Cons: "\<forall>y\<in>set xs. x \<le> y \<Longrightarrow> sorted xs \<Longrightarrow> sorted (x # xs)"
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
   282
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
   283
lemma sorted_single [iff]:
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
   284
  "sorted [x]"
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
   285
  by (rule sorted.Cons) auto
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
   286
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
   287
lemma sorted_many:
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
   288
  "x \<le> y \<Longrightarrow> sorted (y # zs) \<Longrightarrow> sorted (x # y # zs)"
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
   289
  by (rule sorted.Cons) (cases "y # zs" rule: sorted.cases, auto)
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
   290
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
   291
lemma sorted_many_eq [simp, code]:
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
   292
  "sorted (x # y # zs) \<longleftrightarrow> x \<le> y \<and> sorted (y # zs)"
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
   293
  by (auto intro: sorted_many elim: sorted.cases)
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
   294
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
   295
lemma [code]:
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
   296
  "sorted [] \<longleftrightarrow> True"
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
   297
  "sorted [x] \<longleftrightarrow> True"
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
   298
  by simp_all
24697
b37d3980da3c fixed haftmann bug
nipkow
parents: 24657
diff changeset
   299
33639
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
   300
primrec insort_key :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b \<Rightarrow> 'b list \<Rightarrow> 'b list" where
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
   301
"insort_key f x [] = [x]" |
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
   302
"insort_key f x (y#ys) = (if f x \<le> f y then (x#y#ys) else y#(insort_key f x ys))"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
   303
35195
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
   304
definition sort_key :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b list \<Rightarrow> 'b list" where
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
   305
"sort_key f xs = foldr (insort_key f) xs []"
33639
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
   306
40210
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
   307
definition insort_insert_key :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b \<Rightarrow> 'b list \<Rightarrow> 'b list" where
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
   308
  "insort_insert_key f x xs = (if f x \<in> f ` set xs then xs else insort_key f x xs)"
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
   309
33639
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
   310
abbreviation "sort \<equiv> sort_key (\<lambda>x. x)"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
   311
abbreviation "insort \<equiv> insort_key (\<lambda>x. x)"
40210
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
   312
abbreviation "insort_insert \<equiv> insort_insert_key (\<lambda>x. x)"
35608
db4045d1406e added insort_insert
haftmann
parents: 35603
diff changeset
   313
25221
5ded95dda5df append/member: more light-weight way to declare authentic syntax;
wenzelm
parents: 25215
diff changeset
   314
end
5ded95dda5df append/member: more light-weight way to declare authentic syntax;
wenzelm
parents: 25215
diff changeset
   315
24616
fac3dd4ade83 sorting
nipkow
parents: 24566
diff changeset
   316
23388
77645da0db85 tuned proofs: avoid implicit prems;
wenzelm
parents: 23279
diff changeset
   317
subsubsection {* List comprehension *}
23192
ec73b9707d48 Moved list comprehension into List
nipkow
parents: 23096
diff changeset
   318
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
   319
text{* Input syntax for Haskell-like list comprehension notation.
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
   320
Typical example: @{text"[(x,y). x \<leftarrow> xs, y \<leftarrow> ys, x \<noteq> y]"},
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
   321
the list of all pairs of distinct elements from @{text xs} and @{text ys}.
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
   322
The syntax is as in Haskell, except that @{text"|"} becomes a dot
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
   323
(like in Isabelle's set comprehension): @{text"[e. x \<leftarrow> xs, \<dots>]"} rather than
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
   324
\verb![e| x <- xs, ...]!.
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
   325
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
   326
The qualifiers after the dot are
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
   327
\begin{description}
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
   328
\item[generators] @{text"p \<leftarrow> xs"},
24476
f7ad9fbbeeaa turned list comprehension translations into ML to optimize base case
nipkow
parents: 24471
diff changeset
   329
 where @{text p} is a pattern and @{text xs} an expression of list type, or
f7ad9fbbeeaa turned list comprehension translations into ML to optimize base case
nipkow
parents: 24471
diff changeset
   330
\item[guards] @{text"b"}, where @{text b} is a boolean expression.
f7ad9fbbeeaa turned list comprehension translations into ML to optimize base case
nipkow
parents: 24471
diff changeset
   331
%\item[local bindings] @ {text"let x = e"}.
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
   332
\end{description}
23240
7077dc80a14b tuned list comprehension
nipkow
parents: 23235
diff changeset
   333
24476
f7ad9fbbeeaa turned list comprehension translations into ML to optimize base case
nipkow
parents: 24471
diff changeset
   334
Just like in Haskell, list comprehension is just a shorthand. To avoid
f7ad9fbbeeaa turned list comprehension translations into ML to optimize base case
nipkow
parents: 24471
diff changeset
   335
misunderstandings, the translation into desugared form is not reversed
f7ad9fbbeeaa turned list comprehension translations into ML to optimize base case
nipkow
parents: 24471
diff changeset
   336
upon output. Note that the translation of @{text"[e. x \<leftarrow> xs]"} is
f7ad9fbbeeaa turned list comprehension translations into ML to optimize base case
nipkow
parents: 24471
diff changeset
   337
optmized to @{term"map (%x. e) xs"}.
23240
7077dc80a14b tuned list comprehension
nipkow
parents: 23235
diff changeset
   338
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
   339
It is easy to write short list comprehensions which stand for complex
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
   340
expressions. During proofs, they may become unreadable (and
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
   341
mangled). In such cases it can be advisable to introduce separate
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
   342
definitions for the list comprehensions in question.  *}
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
   343
42144
15218eb98fd7 list comprehension: strip positions where the translation cannot handle them right now;
wenzelm
parents: 42057
diff changeset
   344
nonterminal lc_gen and lc_qual and lc_quals
23192
ec73b9707d48 Moved list comprehension into List
nipkow
parents: 23096
diff changeset
   345
ec73b9707d48 Moved list comprehension into List
nipkow
parents: 23096
diff changeset
   346
syntax
23240
7077dc80a14b tuned list comprehension
nipkow
parents: 23235
diff changeset
   347
"_listcompr" :: "'a \<Rightarrow> lc_qual \<Rightarrow> lc_quals \<Rightarrow> 'a list"  ("[_ . __")
42144
15218eb98fd7 list comprehension: strip positions where the translation cannot handle them right now;
wenzelm
parents: 42057
diff changeset
   348
"_lc_gen" :: "lc_gen \<Rightarrow> 'a list \<Rightarrow> lc_qual" ("_ <- _")
23240
7077dc80a14b tuned list comprehension
nipkow
parents: 23235
diff changeset
   349
"_lc_test" :: "bool \<Rightarrow> lc_qual" ("_")
24476
f7ad9fbbeeaa turned list comprehension translations into ML to optimize base case
nipkow
parents: 24471
diff changeset
   350
(*"_lc_let" :: "letbinds => lc_qual"  ("let _")*)
23240
7077dc80a14b tuned list comprehension
nipkow
parents: 23235
diff changeset
   351
"_lc_end" :: "lc_quals" ("]")
7077dc80a14b tuned list comprehension
nipkow
parents: 23235
diff changeset
   352
"_lc_quals" :: "lc_qual \<Rightarrow> lc_quals \<Rightarrow> lc_quals" (", __")
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
   353
"_lc_abs" :: "'a => 'b list => 'b list"
42144
15218eb98fd7 list comprehension: strip positions where the translation cannot handle them right now;
wenzelm
parents: 42057
diff changeset
   354
"_strip_positions" :: "'a \<Rightarrow> lc_gen"  ("_")
23192
ec73b9707d48 Moved list comprehension into List
nipkow
parents: 23096
diff changeset
   355
24476
f7ad9fbbeeaa turned list comprehension translations into ML to optimize base case
nipkow
parents: 24471
diff changeset
   356
(* These are easier than ML code but cannot express the optimized
f7ad9fbbeeaa turned list comprehension translations into ML to optimize base case
nipkow
parents: 24471
diff changeset
   357
   translation of [e. p<-xs]
23192
ec73b9707d48 Moved list comprehension into List
nipkow
parents: 23096
diff changeset
   358
translations
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
   359
"[e. p<-xs]" => "concat(map (_lc_abs p [e]) xs)"
23240
7077dc80a14b tuned list comprehension
nipkow
parents: 23235
diff changeset
   360
"_listcompr e (_lc_gen p xs) (_lc_quals Q Qs)"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
   361
 => "concat (map (_lc_abs p (_listcompr e Q Qs)) xs)"
23240
7077dc80a14b tuned list comprehension
nipkow
parents: 23235
diff changeset
   362
"[e. P]" => "if P then [e] else []"
7077dc80a14b tuned list comprehension
nipkow
parents: 23235
diff changeset
   363
"_listcompr e (_lc_test P) (_lc_quals Q Qs)"
7077dc80a14b tuned list comprehension
nipkow
parents: 23235
diff changeset
   364
 => "if P then (_listcompr e Q Qs) else []"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
   365
"_listcompr e (_lc_let b) (_lc_quals Q Qs)"
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
   366
 => "_Let b (_listcompr e Q Qs)"
24476
f7ad9fbbeeaa turned list comprehension translations into ML to optimize base case
nipkow
parents: 24471
diff changeset
   367
*)
23240
7077dc80a14b tuned list comprehension
nipkow
parents: 23235
diff changeset
   368
23279
e39dd93161d9 tuned list comprehension, changed filter syntax from : to <-
nipkow
parents: 23246
diff changeset
   369
syntax (xsymbols)
42144
15218eb98fd7 list comprehension: strip positions where the translation cannot handle them right now;
wenzelm
parents: 42057
diff changeset
   370
"_lc_gen" :: "lc_gen \<Rightarrow> 'a list \<Rightarrow> lc_qual" ("_ \<leftarrow> _")
23279
e39dd93161d9 tuned list comprehension, changed filter syntax from : to <-
nipkow
parents: 23246
diff changeset
   371
syntax (HTML output)
42144
15218eb98fd7 list comprehension: strip positions where the translation cannot handle them right now;
wenzelm
parents: 42057
diff changeset
   372
"_lc_gen" :: "lc_gen \<Rightarrow> 'a list \<Rightarrow> lc_qual" ("_ \<leftarrow> _")
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
   373
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
   374
parse_translation (advanced) {*
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
   375
let
35256
b73ae1a8fe7e adapted to authentic syntax;
wenzelm
parents: 35248
diff changeset
   376
  val NilC = Syntax.const @{const_syntax Nil};
b73ae1a8fe7e adapted to authentic syntax;
wenzelm
parents: 35248
diff changeset
   377
  val ConsC = Syntax.const @{const_syntax Cons};
b73ae1a8fe7e adapted to authentic syntax;
wenzelm
parents: 35248
diff changeset
   378
  val mapC = Syntax.const @{const_syntax map};
b73ae1a8fe7e adapted to authentic syntax;
wenzelm
parents: 35248
diff changeset
   379
  val concatC = Syntax.const @{const_syntax concat};
b73ae1a8fe7e adapted to authentic syntax;
wenzelm
parents: 35248
diff changeset
   380
  val IfC = Syntax.const @{const_syntax If};
35115
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
   381
24476
f7ad9fbbeeaa turned list comprehension translations into ML to optimize base case
nipkow
parents: 24471
diff changeset
   382
  fun singl x = ConsC $ x $ NilC;
f7ad9fbbeeaa turned list comprehension translations into ML to optimize base case
nipkow
parents: 24471
diff changeset
   383
35115
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
   384
  fun pat_tr ctxt p e opti = (* %x. case x of p => e | _ => [] *)
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
   385
    let
43324
2b47822868e4 discontinued Name.variant to emphasize that this is old-style / indirect;
wenzelm
parents: 42871
diff changeset
   386
      (* FIXME proper name context!? *)
2b47822868e4 discontinued Name.variant to emphasize that this is old-style / indirect;
wenzelm
parents: 42871
diff changeset
   387
      val x = Free (singleton (Name.variant_list (fold Term.add_free_names [p, e] [])) "x", dummyT);
24476
f7ad9fbbeeaa turned list comprehension translations into ML to optimize base case
nipkow
parents: 24471
diff changeset
   388
      val e = if opti then singl e else e;
42264
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents: 42167
diff changeset
   389
      val case1 = Syntax.const @{syntax_const "_case1"} $ Term_Position.strip_positions p $ e;
35256
b73ae1a8fe7e adapted to authentic syntax;
wenzelm
parents: 35248
diff changeset
   390
      val case2 =
b73ae1a8fe7e adapted to authentic syntax;
wenzelm
parents: 35248
diff changeset
   391
        Syntax.const @{syntax_const "_case1"} $
b73ae1a8fe7e adapted to authentic syntax;
wenzelm
parents: 35248
diff changeset
   392
          Syntax.const @{const_syntax dummy_pattern} $ NilC;
35115
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
   393
      val cs = Syntax.const @{syntax_const "_case2"} $ case1 $ case2;
43580
023a1d1f97bd new Datatype.info_of_constr with strict behaviour wrt. to overloaded constructors -- side effect: function package correctly identifies 0::int as a non-constructor;
krauss
parents: 43324
diff changeset
   394
      val ft = Datatype_Case.case_tr false Datatype.info_of_constr_permissive ctxt [x, cs];
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
   395
    in lambda x ft end;
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
   396
35256
b73ae1a8fe7e adapted to authentic syntax;
wenzelm
parents: 35248
diff changeset
   397
  fun abs_tr ctxt (p as Free (s, T)) e opti =
35115
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
   398
        let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   399
          val thy = Proof_Context.theory_of ctxt;
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   400
          val s' = Proof_Context.intern_const ctxt s;
35115
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
   401
        in
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
   402
          if Sign.declared_const thy s'
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
   403
          then (pat_tr ctxt p e opti, false)
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
   404
          else (lambda p e, true)
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
   405
        end
24476
f7ad9fbbeeaa turned list comprehension translations into ML to optimize base case
nipkow
parents: 24471
diff changeset
   406
    | abs_tr ctxt p e opti = (pat_tr ctxt p e opti, false);
f7ad9fbbeeaa turned list comprehension translations into ML to optimize base case
nipkow
parents: 24471
diff changeset
   407
35115
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
   408
  fun lc_tr ctxt [e, Const (@{syntax_const "_lc_test"}, _) $ b, qs] =
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
   409
        let
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
   410
          val res =
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
   411
            (case qs of
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
   412
              Const (@{syntax_const "_lc_end"}, _) => singl e
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
   413
            | Const (@{syntax_const "_lc_quals"}, _) $ q $ qs => lc_tr ctxt [e, q, qs]);
24476
f7ad9fbbeeaa turned list comprehension translations into ML to optimize base case
nipkow
parents: 24471
diff changeset
   414
        in IfC $ b $ res $ NilC end
35115
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
   415
    | lc_tr ctxt
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
   416
          [e, Const (@{syntax_const "_lc_gen"}, _) $ p $ es,
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
   417
            Const(@{syntax_const "_lc_end"}, _)] =
24476
f7ad9fbbeeaa turned list comprehension translations into ML to optimize base case
nipkow
parents: 24471
diff changeset
   418
        (case abs_tr ctxt p e true of
35115
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
   419
          (f, true) => mapC $ f $ es
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
   420
        | (f, false) => concatC $ (mapC $ f $ es))
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
   421
    | lc_tr ctxt
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
   422
          [e, Const (@{syntax_const "_lc_gen"}, _) $ p $ es,
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
   423
            Const (@{syntax_const "_lc_quals"}, _) $ q $ qs] =
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
   424
        let val e' = lc_tr ctxt [e, q, qs];
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
   425
        in concatC $ (mapC $ (fst (abs_tr ctxt p e' false)) $ es) end;
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
   426
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
   427
in [(@{syntax_const "_listcompr"}, lc_tr)] end
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
   428
*}
23279
e39dd93161d9 tuned list comprehension, changed filter syntax from : to <-
nipkow
parents: 23246
diff changeset
   429
42167
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   430
ML {*
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   431
  let
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   432
    val read = Syntax.read_term @{context};
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   433
    fun check s1 s2 = read s1 aconv read s2 orelse error ("Check failed: " ^ quote s1);
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   434
  in
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   435
    check "[(x,y,z). b]" "if b then [(x, y, z)] else []";
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   436
    check "[(x,y,z). x\<leftarrow>xs]" "map (\<lambda>x. (x, y, z)) xs";
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   437
    check "[e x y. x\<leftarrow>xs, y\<leftarrow>ys]" "concat (map (\<lambda>x. map (\<lambda>y. e x y) ys) xs)";
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   438
    check "[(x,y,z). x<a, x>b]" "if x < a then if b < x then [(x, y, z)] else [] else []";
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   439
    check "[(x,y,z). x\<leftarrow>xs, x>b]" "concat (map (\<lambda>x. if b < x then [(x, y, z)] else []) xs)";
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   440
    check "[(x,y,z). x<a, x\<leftarrow>xs]" "if x < a then map (\<lambda>x. (x, y, z)) xs else []";
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   441
    check "[(x,y). Cons True x \<leftarrow> xs]"
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   442
      "concat (map (\<lambda>xa. case xa of [] \<Rightarrow> [] | True # x \<Rightarrow> [(x, y)] | False # x \<Rightarrow> []) xs)";
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   443
    check "[(x,y,z). Cons x [] \<leftarrow> xs]"
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   444
      "concat (map (\<lambda>xa. case xa of [] \<Rightarrow> [] | [x] \<Rightarrow> [(x, y, z)] | x # aa # lista \<Rightarrow> []) xs)";
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   445
    check "[(x,y,z). x<a, x>b, x=d]"
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   446
      "if x < a then if b < x then if x = d then [(x, y, z)] else [] else [] else []";
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   447
    check "[(x,y,z). x<a, x>b, y\<leftarrow>ys]"
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   448
      "if x < a then if b < x then map (\<lambda>y. (x, y, z)) ys else [] else []";
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   449
    check "[(x,y,z). x<a, x\<leftarrow>xs,y>b]"
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   450
      "if x < a then concat (map (\<lambda>x. if b < y then [(x, y, z)] else []) xs) else []";
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   451
    check "[(x,y,z). x<a, x\<leftarrow>xs, y\<leftarrow>ys]"
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   452
      "if x < a then concat (map (\<lambda>x. map (\<lambda>y. (x, y, z)) ys) xs) else []";
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   453
    check "[(x,y,z). x\<leftarrow>xs, x>b, y<a]"
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   454
      "concat (map (\<lambda>x. if b < x then if y < a then [(x, y, z)] else [] else []) xs)";
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   455
    check "[(x,y,z). x\<leftarrow>xs, x>b, y\<leftarrow>ys]"
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   456
      "concat (map (\<lambda>x. if b < x then map (\<lambda>y. (x, y, z)) ys else []) xs)";
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   457
    check "[(x,y,z). x\<leftarrow>xs, y\<leftarrow>ys,y>x]"
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   458
      "concat (map (\<lambda>x. concat (map (\<lambda>y. if x < y then [(x, y, z)] else []) ys)) xs)";
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   459
    check "[(x,y,z). x\<leftarrow>xs, y\<leftarrow>ys,z\<leftarrow>zs]"
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   460
      "concat (map (\<lambda>x. concat (map (\<lambda>y. map (\<lambda>z. (x, y, z)) zs) ys)) xs)"
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   461
  end;
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   462
*}
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   463
35115
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
   464
(*
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
   465
term "[(x,y). x\<leftarrow>xs, let xx = x+x, y\<leftarrow>ys, y \<noteq> xx]"
23192
ec73b9707d48 Moved list comprehension into List
nipkow
parents: 23096
diff changeset
   466
*)
ec73b9707d48 Moved list comprehension into List
nipkow
parents: 23096
diff changeset
   467
42167
7d8cb105373c actually check list comprehension examples;
wenzelm
parents: 42144
diff changeset
   468
41463
edbf0a86fb1c adding simproc to rewrite list comprehensions to set comprehensions; adopting proofs
bulwahn
parents: 41372
diff changeset
   469
use "Tools/list_to_set_comprehension.ML"
edbf0a86fb1c adding simproc to rewrite list comprehensions to set comprehensions; adopting proofs
bulwahn
parents: 41372
diff changeset
   470
edbf0a86fb1c adding simproc to rewrite list comprehensions to set comprehensions; adopting proofs
bulwahn
parents: 41372
diff changeset
   471
simproc_setup list_to_set_comprehension ("set xs") = {* K List_to_Set_Comprehension.simproc *}
edbf0a86fb1c adding simproc to rewrite list comprehensions to set comprehensions; adopting proofs
bulwahn
parents: 41372
diff changeset
   472
35115
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
   473
21061
580dfc999ef6 added normal post setup; cleaned up "execution" constants
haftmann
parents: 21046
diff changeset
   474
subsubsection {* @{const Nil} and @{const Cons} *}
580dfc999ef6 added normal post setup; cleaned up "execution" constants
haftmann
parents: 21046
diff changeset
   475
580dfc999ef6 added normal post setup; cleaned up "execution" constants
haftmann
parents: 21046
diff changeset
   476
lemma not_Cons_self [simp]:
580dfc999ef6 added normal post setup; cleaned up "execution" constants
haftmann
parents: 21046
diff changeset
   477
  "xs \<noteq> x # xs"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   478
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   479
41697
19890332febc explicit is better than implicit;
wenzelm
parents: 41505
diff changeset
   480
lemma not_Cons_self2 [simp]:
19890332febc explicit is better than implicit;
wenzelm
parents: 41505
diff changeset
   481
  "x # xs \<noteq> xs"
19890332febc explicit is better than implicit;
wenzelm
parents: 41505
diff changeset
   482
by (rule not_Cons_self [symmetric])
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   483
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   484
lemma neq_Nil_conv: "(xs \<noteq> []) = (\<exists>y ys. xs = y # ys)"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   485
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   486
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   487
lemma length_induct:
21061
580dfc999ef6 added normal post setup; cleaned up "execution" constants
haftmann
parents: 21046
diff changeset
   488
  "(\<And>xs. \<forall>ys. length ys < length xs \<longrightarrow> P ys \<Longrightarrow> P xs) \<Longrightarrow> P xs"
17589
58eeffd73be1 renamed rules to iprover
nipkow
parents: 17501
diff changeset
   489
by (rule measure_induct [of length]) iprover
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   490
37289
881fa5012451 induction over non-empty lists
haftmann
parents: 37123
diff changeset
   491
lemma list_nonempty_induct [consumes 1, case_names single cons]:
881fa5012451 induction over non-empty lists
haftmann
parents: 37123
diff changeset
   492
  assumes "xs \<noteq> []"
881fa5012451 induction over non-empty lists
haftmann
parents: 37123
diff changeset
   493
  assumes single: "\<And>x. P [x]"
881fa5012451 induction over non-empty lists
haftmann
parents: 37123
diff changeset
   494
  assumes cons: "\<And>x xs. xs \<noteq> [] \<Longrightarrow> P xs \<Longrightarrow> P (x # xs)"
881fa5012451 induction over non-empty lists
haftmann
parents: 37123
diff changeset
   495
  shows "P xs"
881fa5012451 induction over non-empty lists
haftmann
parents: 37123
diff changeset
   496
using `xs \<noteq> []` proof (induct xs)
881fa5012451 induction over non-empty lists
haftmann
parents: 37123
diff changeset
   497
  case Nil then show ?case by simp
881fa5012451 induction over non-empty lists
haftmann
parents: 37123
diff changeset
   498
next
881fa5012451 induction over non-empty lists
haftmann
parents: 37123
diff changeset
   499
  case (Cons x xs) show ?case proof (cases xs)
881fa5012451 induction over non-empty lists
haftmann
parents: 37123
diff changeset
   500
    case Nil with single show ?thesis by simp
881fa5012451 induction over non-empty lists
haftmann
parents: 37123
diff changeset
   501
  next
881fa5012451 induction over non-empty lists
haftmann
parents: 37123
diff changeset
   502
    case Cons then have "xs \<noteq> []" by simp
881fa5012451 induction over non-empty lists
haftmann
parents: 37123
diff changeset
   503
    moreover with Cons.hyps have "P xs" .
881fa5012451 induction over non-empty lists
haftmann
parents: 37123
diff changeset
   504
    ultimately show ?thesis by (rule cons)
881fa5012451 induction over non-empty lists
haftmann
parents: 37123
diff changeset
   505
  qed
881fa5012451 induction over non-empty lists
haftmann
parents: 37123
diff changeset
   506
qed
881fa5012451 induction over non-empty lists
haftmann
parents: 37123
diff changeset
   507
45714
ad4242285560 cardinality of sets of lists
hoelzl
parents: 45607
diff changeset
   508
lemma inj_split_Cons: "inj_on (\<lambda>(xs, n). n#xs) X"
ad4242285560 cardinality of sets of lists
hoelzl
parents: 45607
diff changeset
   509
  by (auto intro!: inj_onI)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   510
21061
580dfc999ef6 added normal post setup; cleaned up "execution" constants
haftmann
parents: 21046
diff changeset
   511
subsubsection {* @{const length} *}
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   512
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   513
text {*
21061
580dfc999ef6 added normal post setup; cleaned up "execution" constants
haftmann
parents: 21046
diff changeset
   514
  Needs to come before @{text "@"} because of theorem @{text
580dfc999ef6 added normal post setup; cleaned up "execution" constants
haftmann
parents: 21046
diff changeset
   515
  append_eq_append_conv}.
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   516
*}
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   517
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   518
lemma length_append [simp]: "length (xs @ ys) = length xs + length ys"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   519
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   520
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   521
lemma length_map [simp]: "length (map f xs) = length xs"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   522
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   523
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   524
lemma length_rev [simp]: "length (rev xs) = length xs"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   525
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   526
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   527
lemma length_tl [simp]: "length (tl xs) = length xs - 1"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   528
by (cases xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   529
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   530
lemma length_0_conv [iff]: "(length xs = 0) = (xs = [])"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   531
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   532
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   533
lemma length_greater_0_conv [iff]: "(0 < length xs) = (xs \<noteq> [])"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   534
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   535
23479
10adbdcdc65b new lemmas
nipkow
parents: 23388
diff changeset
   536
lemma length_pos_if_in_set: "x : set xs \<Longrightarrow> length xs > 0"
10adbdcdc65b new lemmas
nipkow
parents: 23388
diff changeset
   537
by auto
10adbdcdc65b new lemmas
nipkow
parents: 23388
diff changeset
   538
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   539
lemma length_Suc_conv:
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   540
"(length xs = Suc n) = (\<exists>y ys. xs = y # ys \<and> length ys = n)"
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   541
by (induct xs) auto
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   542
14025
d9b155757dc8 *** empty log message ***
nipkow
parents: 13913
diff changeset
   543
lemma Suc_length_conv:
d9b155757dc8 *** empty log message ***
nipkow
parents: 13913
diff changeset
   544
"(Suc n = length xs) = (\<exists>y ys. xs = y # ys \<and> length ys = n)"
14208
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
   545
apply (induct xs, simp, simp)
14025
d9b155757dc8 *** empty log message ***
nipkow
parents: 13913
diff changeset
   546
apply blast
d9b155757dc8 *** empty log message ***
nipkow
parents: 13913
diff changeset
   547
done
d9b155757dc8 *** empty log message ***
nipkow
parents: 13913
diff changeset
   548
25221
5ded95dda5df append/member: more light-weight way to declare authentic syntax;
wenzelm
parents: 25215
diff changeset
   549
lemma impossible_Cons: "length xs <= length ys ==> xs = x # ys = False"
5ded95dda5df append/member: more light-weight way to declare authentic syntax;
wenzelm
parents: 25215
diff changeset
   550
  by (induct xs) auto
5ded95dda5df append/member: more light-weight way to declare authentic syntax;
wenzelm
parents: 25215
diff changeset
   551
26442
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
   552
lemma list_induct2 [consumes 1, case_names Nil Cons]:
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
   553
  "length xs = length ys \<Longrightarrow> P [] [] \<Longrightarrow>
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
   554
   (\<And>x xs y ys. length xs = length ys \<Longrightarrow> P xs ys \<Longrightarrow> P (x#xs) (y#ys))
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
   555
   \<Longrightarrow> P xs ys"
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
   556
proof (induct xs arbitrary: ys)
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
   557
  case Nil then show ?case by simp
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
   558
next
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
   559
  case (Cons x xs ys) then show ?case by (cases ys) simp_all
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
   560
qed
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
   561
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
   562
lemma list_induct3 [consumes 2, case_names Nil Cons]:
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
   563
  "length xs = length ys \<Longrightarrow> length ys = length zs \<Longrightarrow> P [] [] [] \<Longrightarrow>
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
   564
   (\<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))
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
   565
   \<Longrightarrow> P xs ys zs"
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
   566
proof (induct xs arbitrary: ys zs)
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
   567
  case Nil then show ?case by simp
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
   568
next
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
   569
  case (Cons x xs ys zs) then show ?case by (cases ys, simp_all)
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
   570
    (cases zs, simp_all)
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
   571
qed
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   572
36154
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35828
diff changeset
   573
lemma list_induct4 [consumes 3, case_names Nil Cons]:
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35828
diff changeset
   574
  "length xs = length ys \<Longrightarrow> length ys = length zs \<Longrightarrow> length zs = length ws \<Longrightarrow>
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35828
diff changeset
   575
   P [] [] [] [] \<Longrightarrow> (\<And>x xs y ys z zs w ws. length xs = length ys \<Longrightarrow>
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35828
diff changeset
   576
   length ys = length zs \<Longrightarrow> length zs = length ws \<Longrightarrow> P xs ys zs ws \<Longrightarrow>
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35828
diff changeset
   577
   P (x#xs) (y#ys) (z#zs) (w#ws)) \<Longrightarrow> P xs ys zs ws"
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35828
diff changeset
   578
proof (induct xs arbitrary: ys zs ws)
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35828
diff changeset
   579
  case Nil then show ?case by simp
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35828
diff changeset
   580
next
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35828
diff changeset
   581
  case (Cons x xs ys zs ws) then show ?case by ((cases ys, simp_all), (cases zs,simp_all)) (cases ws, simp_all)
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35828
diff changeset
   582
qed
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35828
diff changeset
   583
22493
db930e490fe5 added another rule for simultaneous induction, and lemmas for zip
krauss
parents: 22422
diff changeset
   584
lemma list_induct2': 
db930e490fe5 added another rule for simultaneous induction, and lemmas for zip
krauss
parents: 22422
diff changeset
   585
  "\<lbrakk> P [] [];
db930e490fe5 added another rule for simultaneous induction, and lemmas for zip
krauss
parents: 22422
diff changeset
   586
  \<And>x xs. P (x#xs) [];
db930e490fe5 added another rule for simultaneous induction, and lemmas for zip
krauss
parents: 22422
diff changeset
   587
  \<And>y ys. P [] (y#ys);
db930e490fe5 added another rule for simultaneous induction, and lemmas for zip
krauss
parents: 22422
diff changeset
   588
   \<And>x xs y ys. P xs ys  \<Longrightarrow> P (x#xs) (y#ys) \<rbrakk>
db930e490fe5 added another rule for simultaneous induction, and lemmas for zip
krauss
parents: 22422
diff changeset
   589
 \<Longrightarrow> P xs ys"
db930e490fe5 added another rule for simultaneous induction, and lemmas for zip
krauss
parents: 22422
diff changeset
   590
by (induct xs arbitrary: ys) (case_tac x, auto)+
db930e490fe5 added another rule for simultaneous induction, and lemmas for zip
krauss
parents: 22422
diff changeset
   591
22143
cf58486ca11b Added simproc list_neq (prompted by an application)
nipkow
parents: 21911
diff changeset
   592
lemma neq_if_length_neq: "length xs \<noteq> length ys \<Longrightarrow> (xs = ys) == False"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
   593
by (rule Eq_FalseI) auto
24037
0a41d2ebc0cd proper simproc_setup for "list_neq";
wenzelm
parents: 23983
diff changeset
   594
0a41d2ebc0cd proper simproc_setup for "list_neq";
wenzelm
parents: 23983
diff changeset
   595
simproc_setup list_neq ("(xs::'a list) = ys") = {*
22143
cf58486ca11b Added simproc list_neq (prompted by an application)
nipkow
parents: 21911
diff changeset
   596
(*
cf58486ca11b Added simproc list_neq (prompted by an application)
nipkow
parents: 21911
diff changeset
   597
Reduces xs=ys to False if xs and ys cannot be of the same length.
cf58486ca11b Added simproc list_neq (prompted by an application)
nipkow
parents: 21911
diff changeset
   598
This is the case if the atomic sublists of one are a submultiset
cf58486ca11b Added simproc list_neq (prompted by an application)
nipkow
parents: 21911
diff changeset
   599
of those of the other list and there are fewer Cons's in one than the other.
cf58486ca11b Added simproc list_neq (prompted by an application)
nipkow
parents: 21911
diff changeset
   600
*)
24037
0a41d2ebc0cd proper simproc_setup for "list_neq";
wenzelm
parents: 23983
diff changeset
   601
0a41d2ebc0cd proper simproc_setup for "list_neq";
wenzelm
parents: 23983
diff changeset
   602
let
22143
cf58486ca11b Added simproc list_neq (prompted by an application)
nipkow
parents: 21911
diff changeset
   603
29856
984191be0357 const_name antiquotations
huffman
parents: 29829
diff changeset
   604
fun len (Const(@{const_name Nil},_)) acc = acc
984191be0357 const_name antiquotations
huffman
parents: 29829
diff changeset
   605
  | len (Const(@{const_name Cons},_) $ _ $ xs) (ts,n) = len xs (ts,n+1)
984191be0357 const_name antiquotations
huffman
parents: 29829
diff changeset
   606
  | len (Const(@{const_name append},_) $ xs $ ys) acc = len xs (len ys acc)
984191be0357 const_name antiquotations
huffman
parents: 29829
diff changeset
   607
  | len (Const(@{const_name rev},_) $ xs) acc = len xs acc
984191be0357 const_name antiquotations
huffman
parents: 29829
diff changeset
   608
  | len (Const(@{const_name map},_) $ _ $ xs) acc = len xs acc
22143
cf58486ca11b Added simproc list_neq (prompted by an application)
nipkow
parents: 21911
diff changeset
   609
  | len t (ts,n) = (t::ts,n);
cf58486ca11b Added simproc list_neq (prompted by an application)
nipkow
parents: 21911
diff changeset
   610
24037
0a41d2ebc0cd proper simproc_setup for "list_neq";
wenzelm
parents: 23983
diff changeset
   611
fun list_neq _ ss ct =
22143
cf58486ca11b Added simproc list_neq (prompted by an application)
nipkow
parents: 21911
diff changeset
   612
  let
24037
0a41d2ebc0cd proper simproc_setup for "list_neq";
wenzelm
parents: 23983
diff changeset
   613
    val (Const(_,eqT) $ lhs $ rhs) = Thm.term_of ct;
22143
cf58486ca11b Added simproc list_neq (prompted by an application)
nipkow
parents: 21911
diff changeset
   614
    val (ls,m) = len lhs ([],0) and (rs,n) = len rhs ([],0);
cf58486ca11b Added simproc list_neq (prompted by an application)
nipkow
parents: 21911
diff changeset
   615
    fun prove_neq() =
cf58486ca11b Added simproc list_neq (prompted by an application)
nipkow
parents: 21911
diff changeset
   616
      let
cf58486ca11b Added simproc list_neq (prompted by an application)
nipkow
parents: 21911
diff changeset
   617
        val Type(_,listT::_) = eqT;
22994
02440636214f abstract size function in hologic.ML
haftmann
parents: 22940
diff changeset
   618
        val size = HOLogic.size_const listT;
22143
cf58486ca11b Added simproc list_neq (prompted by an application)
nipkow
parents: 21911
diff changeset
   619
        val eq_len = HOLogic.mk_eq (size $ lhs, size $ rhs);
cf58486ca11b Added simproc list_neq (prompted by an application)
nipkow
parents: 21911
diff changeset
   620
        val neq_len = HOLogic.mk_Trueprop (HOLogic.Not $ eq_len);
cf58486ca11b Added simproc list_neq (prompted by an application)
nipkow
parents: 21911
diff changeset
   621
        val thm = Goal.prove (Simplifier.the_context ss) [] [] neq_len
22633
haftmann
parents: 22551
diff changeset
   622
          (K (simp_tac (Simplifier.inherit_context ss @{simpset}) 1));
haftmann
parents: 22551
diff changeset
   623
      in SOME (thm RS @{thm neq_if_length_neq}) end
22143
cf58486ca11b Added simproc list_neq (prompted by an application)
nipkow
parents: 21911
diff changeset
   624
  in
23214
dc23c062b58c renamed gen_submultiset to submultiset;
wenzelm
parents: 23212
diff changeset
   625
    if m < n andalso submultiset (op aconv) (ls,rs) orelse
dc23c062b58c renamed gen_submultiset to submultiset;
wenzelm
parents: 23212
diff changeset
   626
       n < m andalso submultiset (op aconv) (rs,ls)
22143
cf58486ca11b Added simproc list_neq (prompted by an application)
nipkow
parents: 21911
diff changeset
   627
    then prove_neq() else NONE
cf58486ca11b Added simproc list_neq (prompted by an application)
nipkow
parents: 21911
diff changeset
   628
  end;
24037
0a41d2ebc0cd proper simproc_setup for "list_neq";
wenzelm
parents: 23983
diff changeset
   629
in list_neq end;
22143
cf58486ca11b Added simproc list_neq (prompted by an application)
nipkow
parents: 21911
diff changeset
   630
*}
cf58486ca11b Added simproc list_neq (prompted by an application)
nipkow
parents: 21911
diff changeset
   631
cf58486ca11b Added simproc list_neq (prompted by an application)
nipkow
parents: 21911
diff changeset
   632
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15307
diff changeset
   633
subsubsection {* @{text "@"} -- append *}
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   634
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   635
lemma append_assoc [simp]: "(xs @ ys) @ zs = xs @ (ys @ zs)"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   636
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   637
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   638
lemma append_Nil2 [simp]: "xs @ [] = xs"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   639
by (induct xs) auto
3507
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
   640
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   641
lemma append_is_Nil_conv [iff]: "(xs @ ys = []) = (xs = [] \<and> ys = [])"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   642
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   643
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   644
lemma Nil_is_append_conv [iff]: "([] = xs @ ys) = (xs = [] \<and> ys = [])"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   645
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   646
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   647
lemma append_self_conv [iff]: "(xs @ ys = xs) = (ys = [])"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   648
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   649
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   650
lemma self_append_conv [iff]: "(xs = xs @ ys) = (ys = [])"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   651
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   652
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35827
diff changeset
   653
lemma append_eq_append_conv [simp, no_atp]:
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
   654
 "length xs = length ys \<or> length us = length vs
13883
0451e0fb3f22 Re-structured some proofs in order to get rid of rule_format attribute.
berghofe
parents: 13863
diff changeset
   655
 ==> (xs@us = ys@vs) = (xs=ys \<and> us=vs)"
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
   656
apply (induct xs arbitrary: ys)
14208
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
   657
 apply (case_tac ys, simp, force)
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
   658
apply (case_tac ys, force, simp)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   659
done
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   660
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
   661
lemma append_eq_append_conv2: "(xs @ ys = zs @ ts) =
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
   662
  (EX us. xs = zs @ us & us @ ys = ts | xs @ us = zs & ys = us@ ts)"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
   663
apply (induct xs arbitrary: ys zs ts)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44635
diff changeset
   664
 apply fastforce
14495
e2a1c31cf6d3 Added append_eq_append_conv2
nipkow
parents: 14402
diff changeset
   665
apply(case_tac zs)
e2a1c31cf6d3 Added append_eq_append_conv2
nipkow
parents: 14402
diff changeset
   666
 apply simp
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44635
diff changeset
   667
apply fastforce
14495
e2a1c31cf6d3 Added append_eq_append_conv2
nipkow
parents: 14402
diff changeset
   668
done
e2a1c31cf6d3 Added append_eq_append_conv2
nipkow
parents: 14402
diff changeset
   669
34910
b23bd3ee4813 same_append_eq / append_same_eq are now used for simplifying induction rules.
berghofe
parents: 34064
diff changeset
   670
lemma same_append_eq [iff, induct_simp]: "(xs @ ys = xs @ zs) = (ys = zs)"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   671
by simp
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   672
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   673
lemma append1_eq_conv [iff]: "(xs @ [x] = ys @ [y]) = (xs = ys \<and> x = y)"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   674
by simp
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   675
34910
b23bd3ee4813 same_append_eq / append_same_eq are now used for simplifying induction rules.
berghofe
parents: 34064
diff changeset
   676
lemma append_same_eq [iff, induct_simp]: "(ys @ xs = zs @ xs) = (ys = zs)"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   677
by simp
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   678
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   679
lemma append_self_conv2 [iff]: "(xs @ ys = ys) = (xs = [])"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   680
using append_same_eq [of _ _ "[]"] by auto
3507
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
   681
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   682
lemma self_append_conv2 [iff]: "(ys = xs @ ys) = (xs = [])"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   683
using append_same_eq [of "[]"] by auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   684
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35827
diff changeset
   685
lemma hd_Cons_tl [simp,no_atp]: "xs \<noteq> [] ==> hd xs # tl xs = xs"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   686
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   687
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   688
lemma hd_append: "hd (xs @ ys) = (if xs = [] then hd ys else hd xs)"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   689
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   690
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   691
lemma hd_append2 [simp]: "xs \<noteq> [] ==> hd (xs @ ys) = hd xs"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   692
by (simp add: hd_append split: list.split)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   693
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   694
lemma tl_append: "tl (xs @ ys) = (case xs of [] => tl ys | z#zs => zs @ ys)"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   695
by (simp split: list.split)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   696
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   697
lemma tl_append2 [simp]: "xs \<noteq> [] ==> tl (xs @ ys) = tl xs @ ys"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   698
by (simp add: tl_append split: list.split)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   699
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   700
14300
bf8b8c9425c3 *** empty log message ***
nipkow
parents: 14247
diff changeset
   701
lemma Cons_eq_append_conv: "x#xs = ys@zs =
bf8b8c9425c3 *** empty log message ***
nipkow
parents: 14247
diff changeset
   702
 (ys = [] & x#xs = zs | (EX ys'. x#ys' = ys & xs = ys'@zs))"
bf8b8c9425c3 *** empty log message ***
nipkow
parents: 14247
diff changeset
   703
by(cases ys) auto
bf8b8c9425c3 *** empty log message ***
nipkow
parents: 14247
diff changeset
   704
15281
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
   705
lemma append_eq_Cons_conv: "(ys@zs = x#xs) =
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
   706
 (ys = [] & zs = x#xs | (EX ys'. ys = x#ys' & ys'@zs = xs))"
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
   707
by(cases ys) auto
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
   708
14300
bf8b8c9425c3 *** empty log message ***
nipkow
parents: 14247
diff changeset
   709
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   710
text {* Trivial rules for solving @{text "@"}-equations automatically. *}
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   711
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   712
lemma eq_Nil_appendI: "xs = ys ==> xs = [] @ ys"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   713
by simp
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   714
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   715
lemma Cons_eq_appendI:
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   716
"[| x # xs1 = ys; xs = xs1 @ zs |] ==> x # xs = ys @ zs"
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   717
by (drule sym) simp
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   718
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   719
lemma append_eq_appendI:
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   720
"[| xs @ xs1 = zs; ys = xs1 @ us |] ==> xs @ ys = zs @ us"
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   721
by (drule sym) simp
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   722
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   723
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   724
text {*
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   725
Simplification procedure for all list equalities.
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   726
Currently only tries to rearrange @{text "@"} to see if
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   727
- both lists end in a singleton list,
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   728
- or both lists end in the same list.
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   729
*}
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   730
43594
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   731
simproc_setup list_eq ("(xs::'a list) = ys")  = {*
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 13366
diff changeset
   732
  let
43594
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   733
    fun last (cons as Const (@{const_name Cons}, _) $ _ $ xs) =
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   734
          (case xs of Const (@{const_name Nil}, _) => cons | _ => last xs)
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   735
      | last (Const(@{const_name append},_) $ _ $ ys) = last ys
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   736
      | last t = t;
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   737
    
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   738
    fun list1 (Const(@{const_name Cons},_) $ _ $ Const(@{const_name Nil},_)) = true
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   739
      | list1 _ = false;
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   740
    
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   741
    fun butlast ((cons as Const(@{const_name Cons},_) $ x) $ xs) =
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   742
          (case xs of Const (@{const_name Nil}, _) => xs | _ => cons $ butlast xs)
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   743
      | butlast ((app as Const (@{const_name append}, _) $ xs) $ ys) = app $ butlast ys
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   744
      | butlast xs = Const(@{const_name Nil}, fastype_of xs);
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   745
    
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   746
    val rearr_ss =
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   747
      HOL_basic_ss addsimps [@{thm append_assoc}, @{thm append_Nil}, @{thm append_Cons}];
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   748
    
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   749
    fun list_eq ss (F as (eq as Const(_,eqT)) $ lhs $ rhs) =
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 13366
diff changeset
   750
      let
43594
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   751
        val lastl = last lhs and lastr = last rhs;
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   752
        fun rearr conv =
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   753
          let
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   754
            val lhs1 = butlast lhs and rhs1 = butlast rhs;
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   755
            val Type(_,listT::_) = eqT
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   756
            val appT = [listT,listT] ---> listT
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   757
            val app = Const(@{const_name append},appT)
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   758
            val F2 = eq $ (app$lhs1$lastl) $ (app$rhs1$lastr)
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   759
            val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (F,F2));
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   760
            val thm = Goal.prove (Simplifier.the_context ss) [] [] eq
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   761
              (K (simp_tac (Simplifier.inherit_context ss rearr_ss) 1));
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   762
          in SOME ((conv RS (thm RS trans)) RS eq_reflection) end;
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   763
      in
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   764
        if list1 lastl andalso list1 lastr then rearr @{thm append1_eq_conv}
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   765
        else if lastl aconv lastr then rearr @{thm append_same_eq}
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   766
        else NONE
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   767
      end;
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 43580
diff changeset
   768
  in fn _ => fn ss => fn ct => list_eq ss (term_of ct) end;
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   769
*}
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   770
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   771
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15307
diff changeset
   772
subsubsection {* @{text map} *}
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   773
40210
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
   774
lemma hd_map:
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
   775
  "xs \<noteq> [] \<Longrightarrow> hd (map f xs) = f (hd xs)"
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
   776
  by (cases xs) simp_all
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
   777
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
   778
lemma map_tl:
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
   779
  "map f (tl xs) = tl (map f xs)"
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
   780
  by (cases xs) simp_all
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
   781
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   782
lemma map_ext: "(!!x. x : set xs --> f x = g x) ==> map f xs = map g xs"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   783
by (induct xs) simp_all
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   784
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   785
lemma map_ident [simp]: "map (\<lambda>x. x) = (\<lambda>xs. xs)"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   786
by (rule ext, induct_tac xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   787
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   788
lemma map_append [simp]: "map f (xs @ ys) = map f xs @ map f ys"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   789
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   790
33639
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
   791
lemma map_map [simp]: "map f (map g xs) = map (f \<circ> g) xs"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
   792
by (induct xs) auto
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
   793
35208
2b9bce05e84b added lemma
nipkow
parents: 35195
diff changeset
   794
lemma map_comp_map[simp]: "((map f) o (map g)) = map(f o g)"
2b9bce05e84b added lemma
nipkow
parents: 35195
diff changeset
   795
apply(rule ext)
2b9bce05e84b added lemma
nipkow
parents: 35195
diff changeset
   796
apply(simp)
2b9bce05e84b added lemma
nipkow
parents: 35195
diff changeset
   797
done
2b9bce05e84b added lemma
nipkow
parents: 35195
diff changeset
   798
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   799
lemma rev_map: "rev (map f xs) = map f (rev xs)"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   800
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   801
13737
e564c3d2d174 added a few lemmas
nipkow
parents: 13601
diff changeset
   802
lemma map_eq_conv[simp]: "(map f xs = map g xs) = (!x : set xs. f x = g x)"
e564c3d2d174 added a few lemmas
nipkow
parents: 13601
diff changeset
   803
by (induct xs) auto
e564c3d2d174 added a few lemmas
nipkow
parents: 13601
diff changeset
   804
44013
5cfc1c36ae97 moved recdef package to HOL/Library/Old_Recdef.thy
krauss
parents: 43594
diff changeset
   805
lemma map_cong [fundef_cong]:
40122
1d8ad2ff3e01 dropped (almost) redundant distinct.induct rule; distinct_simps again named distinct.simps
haftmann
parents: 40077
diff changeset
   806
  "xs = ys \<Longrightarrow> (\<And>x. x \<in> set ys \<Longrightarrow> f x = g x) \<Longrightarrow> map f xs = map g ys"
1d8ad2ff3e01 dropped (almost) redundant distinct.induct rule; distinct_simps again named distinct.simps
haftmann
parents: 40077
diff changeset
   807
  by simp
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   808
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   809
lemma map_is_Nil_conv [iff]: "(map f xs = []) = (xs = [])"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   810
by (cases xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   811
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   812
lemma Nil_is_map_conv [iff]: "([] = map f xs) = (xs = [])"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   813
by (cases xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   814
18447
da548623916a removed or modified some instances of [iff]
paulson
parents: 18423
diff changeset
   815
lemma map_eq_Cons_conv:
14025
d9b155757dc8 *** empty log message ***
nipkow
parents: 13913
diff changeset
   816
 "(map f xs = y#ys) = (\<exists>z zs. xs = z#zs \<and> f z = y \<and> map f zs = ys)"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   817
by (cases xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   818
18447
da548623916a removed or modified some instances of [iff]
paulson
parents: 18423
diff changeset
   819
lemma Cons_eq_map_conv:
14025
d9b155757dc8 *** empty log message ***
nipkow
parents: 13913
diff changeset
   820
 "(x#xs = map f ys) = (\<exists>z zs. ys = z#zs \<and> x = f z \<and> xs = map f zs)"
d9b155757dc8 *** empty log message ***
nipkow
parents: 13913
diff changeset
   821
by (cases ys) auto
d9b155757dc8 *** empty log message ***
nipkow
parents: 13913
diff changeset
   822
18447
da548623916a removed or modified some instances of [iff]
paulson
parents: 18423
diff changeset
   823
lemmas map_eq_Cons_D = map_eq_Cons_conv [THEN iffD1]
da548623916a removed or modified some instances of [iff]
paulson
parents: 18423
diff changeset
   824
lemmas Cons_eq_map_D = Cons_eq_map_conv [THEN iffD1]
da548623916a removed or modified some instances of [iff]
paulson
parents: 18423
diff changeset
   825
declare map_eq_Cons_D [dest!]  Cons_eq_map_D [dest!]
da548623916a removed or modified some instances of [iff]
paulson
parents: 18423
diff changeset
   826
14111
993471c762b8 Some new thm (ex_map_conv?)
nipkow
parents: 14099
diff changeset
   827
lemma ex_map_conv:
993471c762b8 Some new thm (ex_map_conv?)
nipkow
parents: 14099
diff changeset
   828
  "(EX xs. ys = map f xs) = (ALL y : set ys. EX x. y = f x)"
18447
da548623916a removed or modified some instances of [iff]
paulson
parents: 18423
diff changeset
   829
by(induct ys, auto simp add: Cons_eq_map_conv)
14111
993471c762b8 Some new thm (ex_map_conv?)
nipkow
parents: 14099
diff changeset
   830
15110
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
   831
lemma map_eq_imp_length_eq:
35510
64d2d54cbf03 Slightly generalised a theorem
paulson
parents: 35296
diff changeset
   832
  assumes "map f xs = map g ys"
26734
a92057c1ee21 dropped some metis calls
haftmann
parents: 26584
diff changeset
   833
  shows "length xs = length ys"
a92057c1ee21 dropped some metis calls
haftmann
parents: 26584
diff changeset
   834
using assms proof (induct ys arbitrary: xs)
a92057c1ee21 dropped some metis calls
haftmann
parents: 26584
diff changeset
   835
  case Nil then show ?case by simp
a92057c1ee21 dropped some metis calls
haftmann
parents: 26584
diff changeset
   836
next
a92057c1ee21 dropped some metis calls
haftmann
parents: 26584
diff changeset
   837
  case (Cons y ys) then obtain z zs where xs: "xs = z # zs" by auto
35510
64d2d54cbf03 Slightly generalised a theorem
paulson
parents: 35296
diff changeset
   838
  from Cons xs have "map f zs = map g ys" by simp
26734
a92057c1ee21 dropped some metis calls
haftmann
parents: 26584
diff changeset
   839
  moreover with Cons have "length zs = length ys" by blast
a92057c1ee21 dropped some metis calls
haftmann
parents: 26584
diff changeset
   840
  with xs show ?case by simp
a92057c1ee21 dropped some metis calls
haftmann
parents: 26584
diff changeset
   841
qed
a92057c1ee21 dropped some metis calls
haftmann
parents: 26584
diff changeset
   842
  
15110
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
   843
lemma map_inj_on:
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
   844
 "[| map f xs = map f ys; inj_on f (set xs Un set ys) |]
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
   845
  ==> xs = ys"
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
   846
apply(frule map_eq_imp_length_eq)
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
   847
apply(rotate_tac -1)
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
   848
apply(induct rule:list_induct2)
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
   849
 apply simp
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
   850
apply(simp)
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
   851
apply (blast intro:sym)
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
   852
done
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
   853
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
   854
lemma inj_on_map_eq_map:
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
   855
 "inj_on f (set xs Un set ys) \<Longrightarrow> (map f xs = map f ys) = (xs = ys)"
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
   856
by(blast dest:map_inj_on)
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
   857
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   858
lemma map_injective:
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
   859
 "map f xs = map f ys ==> inj f ==> xs = ys"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
   860
by (induct ys arbitrary: xs) (auto dest!:injD)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   861
14339
ec575b7bde7a *** empty log message ***
nipkow
parents: 14338
diff changeset
   862
lemma inj_map_eq_map[simp]: "inj f \<Longrightarrow> (map f xs = map f ys) = (xs = ys)"
ec575b7bde7a *** empty log message ***
nipkow
parents: 14338
diff changeset
   863
by(blast dest:map_injective)
ec575b7bde7a *** empty log message ***
nipkow
parents: 14338
diff changeset
   864
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   865
lemma inj_mapI: "inj f ==> inj (map f)"
17589
58eeffd73be1 renamed rules to iprover
nipkow
parents: 17501
diff changeset
   866
by (iprover dest: map_injective injD intro: inj_onI)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   867
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   868
lemma inj_mapD: "inj (map f) ==> inj f"
14208
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
   869
apply (unfold inj_on_def, clarify)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   870
apply (erule_tac x = "[x]" in ballE)
14208
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
   871
 apply (erule_tac x = "[y]" in ballE, simp, blast)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   872
apply blast
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   873
done
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   874
14339
ec575b7bde7a *** empty log message ***
nipkow
parents: 14338
diff changeset
   875
lemma inj_map[iff]: "inj (map f) = inj f"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   876
by (blast dest: inj_mapD intro: inj_mapI)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   877
15303
eedbb8d22ca2 added lemmas
nipkow
parents: 15302
diff changeset
   878
lemma inj_on_mapI: "inj_on f (\<Union>(set ` A)) \<Longrightarrow> inj_on (map f) A"
eedbb8d22ca2 added lemmas
nipkow
parents: 15302
diff changeset
   879
apply(rule inj_onI)
eedbb8d22ca2 added lemmas
nipkow
parents: 15302
diff changeset
   880
apply(erule map_inj_on)
eedbb8d22ca2 added lemmas
nipkow
parents: 15302
diff changeset
   881
apply(blast intro:inj_onI dest:inj_onD)
eedbb8d22ca2 added lemmas
nipkow
parents: 15302
diff changeset
   882
done
eedbb8d22ca2 added lemmas
nipkow
parents: 15302
diff changeset
   883
14343
6bc647f472b9 map_idI
kleing
parents: 14339
diff changeset
   884
lemma map_idI: "(\<And>x. x \<in> set xs \<Longrightarrow> f x = x) \<Longrightarrow> map f xs = xs"
6bc647f472b9 map_idI
kleing
parents: 14339
diff changeset
   885
by (induct xs, auto)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   886
14402
4201e1916482 moved lemmas from MicroJava/Comp/AuxLemmas.thy to List.thy
nipkow
parents: 14395
diff changeset
   887
lemma map_fun_upd [simp]: "y \<notin> set xs \<Longrightarrow> map (f(y:=v)) xs = map f xs"
4201e1916482 moved lemmas from MicroJava/Comp/AuxLemmas.thy to List.thy
nipkow
parents: 14395
diff changeset
   888
by (induct xs) auto
4201e1916482 moved lemmas from MicroJava/Comp/AuxLemmas.thy to List.thy
nipkow
parents: 14395
diff changeset
   889
15110
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
   890
lemma map_fst_zip[simp]:
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
   891
  "length xs = length ys \<Longrightarrow> map fst (zip xs ys) = xs"
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
   892
by (induct rule:list_induct2, simp_all)
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
   893
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
   894
lemma map_snd_zip[simp]:
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
   895
  "length xs = length ys \<Longrightarrow> map snd (zip xs ys) = ys"
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
   896
by (induct rule:list_induct2, simp_all)
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
   897
41505
6d19301074cf "enriched_type" replaces less specific "type_lifting"
haftmann
parents: 41463
diff changeset
   898
enriched_type map: map
41372
551eb49a6e91 tuned type_lifting declarations
haftmann
parents: 41229
diff changeset
   899
  by (simp_all add: fun_eq_iff id_def)
40608
6c28ab8b8166 mapper for list type; map_pair replaces prod_fun
haftmann
parents: 40593
diff changeset
   900
15110
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
   901
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15307
diff changeset
   902
subsubsection {* @{text rev} *}
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   903
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   904
lemma rev_append [simp]: "rev (xs @ ys) = rev ys @ rev xs"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   905
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   906
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   907
lemma rev_rev_ident [simp]: "rev (rev xs) = xs"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   908
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   909
15870
4320bce5873f more on rev
kleing
parents: 15868
diff changeset
   910
lemma rev_swap: "(rev xs = ys) = (xs = rev ys)"
4320bce5873f more on rev
kleing
parents: 15868
diff changeset
   911
by auto
4320bce5873f more on rev
kleing
parents: 15868
diff changeset
   912
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   913
lemma rev_is_Nil_conv [iff]: "(rev xs = []) = (xs = [])"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   914
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   915
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   916
lemma Nil_is_rev_conv [iff]: "([] = rev xs) = (xs = [])"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   917
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   918
15870
4320bce5873f more on rev
kleing
parents: 15868
diff changeset
   919
lemma rev_singleton_conv [simp]: "(rev xs = [x]) = (xs = [x])"
4320bce5873f more on rev
kleing
parents: 15868
diff changeset
   920
by (cases xs) auto
4320bce5873f more on rev
kleing
parents: 15868
diff changeset
   921
4320bce5873f more on rev
kleing
parents: 15868
diff changeset
   922
lemma singleton_rev_conv [simp]: "([x] = rev xs) = (xs = [x])"
4320bce5873f more on rev
kleing
parents: 15868
diff changeset
   923
by (cases xs) auto
4320bce5873f more on rev
kleing
parents: 15868
diff changeset
   924
21061
580dfc999ef6 added normal post setup; cleaned up "execution" constants
haftmann
parents: 21046
diff changeset
   925
lemma rev_is_rev_conv [iff]: "(rev xs = rev ys) = (xs = ys)"
580dfc999ef6 added normal post setup; cleaned up "execution" constants
haftmann
parents: 21046
diff changeset
   926
apply (induct xs arbitrary: ys, force)
14208
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
   927
apply (case_tac ys, simp, force)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   928
done
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   929
15439
71c0f98e31f1 made diff_less a simp rule
nipkow
parents: 15426
diff changeset
   930
lemma inj_on_rev[iff]: "inj_on rev A"
71c0f98e31f1 made diff_less a simp rule
nipkow
parents: 15426
diff changeset
   931
by(simp add:inj_on_def)
71c0f98e31f1 made diff_less a simp rule
nipkow
parents: 15426
diff changeset
   932
13366
114b7c14084a moved stuff from Main.thy;
wenzelm
parents: 13187
diff changeset
   933
lemma rev_induct [case_names Nil snoc]:
114b7c14084a moved stuff from Main.thy;
wenzelm
parents: 13187
diff changeset
   934
  "[| P []; !!x xs. P xs ==> P (xs @ [x]) |] ==> P xs"
15489
d136af442665 Replaced application of subst by simplesubst in proof of rev_induct
berghofe
parents: 15439
diff changeset
   935
apply(simplesubst rev_rev_ident[symmetric])
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   936
apply(rule_tac list = "rev xs" in list.induct, simp_all)
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   937
done
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   938
13366
114b7c14084a moved stuff from Main.thy;
wenzelm
parents: 13187
diff changeset
   939
lemma rev_exhaust [case_names Nil snoc]:
114b7c14084a moved stuff from Main.thy;
wenzelm
parents: 13187
diff changeset
   940
  "(xs = [] ==> P) ==>(!!ys y. xs = ys @ [y] ==> P) ==> P"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   941
by (induct xs rule: rev_induct) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   942
13366
114b7c14084a moved stuff from Main.thy;
wenzelm
parents: 13187
diff changeset
   943
lemmas rev_cases = rev_exhaust
114b7c14084a moved stuff from Main.thy;
wenzelm
parents: 13187
diff changeset
   944
18423
d7859164447f new lemmas
nipkow
parents: 18336
diff changeset
   945
lemma rev_eq_Cons_iff[iff]: "(rev xs = y#ys) = (xs = rev ys @ [y])"
d7859164447f new lemmas
nipkow
parents: 18336
diff changeset
   946
by(rule rev_cases[of xs]) auto
d7859164447f new lemmas
nipkow
parents: 18336
diff changeset
   947
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   948
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15307
diff changeset
   949
subsubsection {* @{text set} *}
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   950
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   951
lemma finite_set [iff]: "finite (set xs)"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   952
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   953
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   954
lemma set_append [simp]: "set (xs @ ys) = (set xs \<union> set ys)"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   955
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   956
17830
695a2365d32f added hd lemma
nipkow
parents: 17765
diff changeset
   957
lemma hd_in_set[simp]: "xs \<noteq> [] \<Longrightarrow> hd xs : set xs"
695a2365d32f added hd lemma
nipkow
parents: 17765
diff changeset
   958
by(cases xs) auto
14099
55d244f3c86d added fold_red, o2l, postfix, some thms
oheimb
parents: 14050
diff changeset
   959
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   960
lemma set_subset_Cons: "set xs \<subseteq> set (x # xs)"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   961
by auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   962
14099
55d244f3c86d added fold_red, o2l, postfix, some thms
oheimb
parents: 14050
diff changeset
   963
lemma set_ConsD: "y \<in> set (x # xs) \<Longrightarrow> y=x \<or> y \<in> set xs" 
55d244f3c86d added fold_red, o2l, postfix, some thms
oheimb
parents: 14050
diff changeset
   964
by auto
55d244f3c86d added fold_red, o2l, postfix, some thms
oheimb
parents: 14050
diff changeset
   965
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   966
lemma set_empty [iff]: "(set xs = {}) = (xs = [])"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   967
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   968
15245
5a21d9a8f14b Added a few lemmas
nipkow
parents: 15236
diff changeset
   969
lemma set_empty2[iff]: "({} = set xs) = (xs = [])"
5a21d9a8f14b Added a few lemmas
nipkow
parents: 15236
diff changeset
   970
by(induct xs) auto
5a21d9a8f14b Added a few lemmas
nipkow
parents: 15236
diff changeset
   971
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   972
lemma set_rev [simp]: "set (rev xs) = set xs"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   973
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   974
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   975
lemma set_map [simp]: "set (map f xs) = f`(set xs)"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   976
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   977
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   978
lemma set_filter [simp]: "set (filter P xs) = {x. x : set xs \<and> P x}"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
   979
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   980
32417
e87d9c78910c tuned code generation for lists
nipkow
parents: 32415
diff changeset
   981
lemma set_upt [simp]: "set[i..<j] = {i..<j}"
41463
edbf0a86fb1c adding simproc to rewrite list comprehensions to set comprehensions; adopting proofs
bulwahn
parents: 41372
diff changeset
   982
by (induct j) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
   983
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
   984
25221
5ded95dda5df append/member: more light-weight way to declare authentic syntax;
wenzelm
parents: 25215
diff changeset
   985
lemma split_list: "x : set xs \<Longrightarrow> \<exists>ys zs. xs = ys @ x # zs"
18049
156bba334c12 A few new lemmas
nipkow
parents: 17956
diff changeset
   986
proof (induct xs)
26073
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
   987
  case Nil thus ?case by simp
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
   988
next
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
   989
  case Cons thus ?case by (auto intro: Cons_eq_appendI)
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
   990
qed
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
   991
26734
a92057c1ee21 dropped some metis calls
haftmann
parents: 26584
diff changeset
   992
lemma in_set_conv_decomp: "x \<in> set xs \<longleftrightarrow> (\<exists>ys zs. xs = ys @ x # zs)"
a92057c1ee21 dropped some metis calls
haftmann
parents: 26584
diff changeset
   993
  by (auto elim: split_list)
26073
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
   994
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
   995
lemma split_list_first: "x : set xs \<Longrightarrow> \<exists>ys zs. xs = ys @ x # zs \<and> x \<notin> set ys"
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
   996
proof (induct xs)
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
   997
  case Nil thus ?case by simp
18049
156bba334c12 A few new lemmas
nipkow
parents: 17956
diff changeset
   998
next
156bba334c12 A few new lemmas
nipkow
parents: 17956
diff changeset
   999
  case (Cons a xs)
156bba334c12 A few new lemmas
nipkow
parents: 17956
diff changeset
  1000
  show ?case
156bba334c12 A few new lemmas
nipkow
parents: 17956
diff changeset
  1001
  proof cases
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44635
diff changeset
  1002
    assume "x = a" thus ?case using Cons by fastforce
18049
156bba334c12 A few new lemmas
nipkow
parents: 17956
diff changeset
  1003
  next
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44635
diff changeset
  1004
    assume "x \<noteq> a" thus ?case using Cons by(fastforce intro!: Cons_eq_appendI)
26073
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1005
  qed
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1006
qed
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1007
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1008
lemma in_set_conv_decomp_first:
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1009
  "(x : set xs) = (\<exists>ys zs. xs = ys @ x # zs \<and> x \<notin> set ys)"
26734
a92057c1ee21 dropped some metis calls
haftmann
parents: 26584
diff changeset
  1010
  by (auto dest!: split_list_first)
26073
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1011
40122
1d8ad2ff3e01 dropped (almost) redundant distinct.induct rule; distinct_simps again named distinct.simps
haftmann
parents: 40077
diff changeset
  1012
lemma split_list_last: "x \<in> set xs \<Longrightarrow> \<exists>ys zs. xs = ys @ x # zs \<and> x \<notin> set zs"
1d8ad2ff3e01 dropped (almost) redundant distinct.induct rule; distinct_simps again named distinct.simps
haftmann
parents: 40077
diff changeset
  1013
proof (induct xs rule: rev_induct)
26073
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1014
  case Nil thus ?case by simp
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1015
next
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1016
  case (snoc a xs)
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1017
  show ?case
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1018
  proof cases
40122
1d8ad2ff3e01 dropped (almost) redundant distinct.induct rule; distinct_simps again named distinct.simps
haftmann
parents: 40077
diff changeset
  1019
    assume "x = a" thus ?case using snoc by (metis List.set.simps(1) emptyE)
26073
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1020
  next
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44635
diff changeset
  1021
    assume "x \<noteq> a" thus ?case using snoc by fastforce
18049
156bba334c12 A few new lemmas
nipkow
parents: 17956
diff changeset
  1022
  qed
156bba334c12 A few new lemmas
nipkow
parents: 17956
diff changeset
  1023
qed
156bba334c12 A few new lemmas
nipkow
parents: 17956
diff changeset
  1024
26073
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1025
lemma in_set_conv_decomp_last:
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1026
  "(x : set xs) = (\<exists>ys zs. xs = ys @ x # zs \<and> x \<notin> set zs)"
26734
a92057c1ee21 dropped some metis calls
haftmann
parents: 26584
diff changeset
  1027
  by (auto dest!: split_list_last)
26073
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1028
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1029
lemma split_list_prop: "\<exists>x \<in> set xs. P x \<Longrightarrow> \<exists>ys x zs. xs = ys @ x # zs & P x"
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1030
proof (induct xs)
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1031
  case Nil thus ?case by simp
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1032
next
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1033
  case Cons thus ?case
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1034
    by(simp add:Bex_def)(metis append_Cons append.simps(1))
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1035
qed
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1036
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1037
lemma split_list_propE:
26734
a92057c1ee21 dropped some metis calls
haftmann
parents: 26584
diff changeset
  1038
  assumes "\<exists>x \<in> set xs. P x"
a92057c1ee21 dropped some metis calls
haftmann
parents: 26584
diff changeset
  1039
  obtains ys x zs where "xs = ys @ x # zs" and "P x"
a92057c1ee21 dropped some metis calls
haftmann
parents: 26584
diff changeset
  1040
using split_list_prop [OF assms] by blast
26073
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1041
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1042
lemma split_list_first_prop:
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1043
  "\<exists>x \<in> set xs. P x \<Longrightarrow>
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1044
   \<exists>ys x zs. xs = ys@x#zs \<and> P x \<and> (\<forall>y \<in> set ys. \<not> P y)"
26734
a92057c1ee21 dropped some metis calls
haftmann
parents: 26584
diff changeset
  1045
proof (induct xs)
26073
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1046
  case Nil thus ?case by simp
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1047
next
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1048
  case (Cons x xs)
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1049
  show ?case
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1050
  proof cases
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1051
    assume "P x"
40122
1d8ad2ff3e01 dropped (almost) redundant distinct.induct rule; distinct_simps again named distinct.simps
haftmann
parents: 40077
diff changeset
  1052
    thus ?thesis by simp (metis Un_upper1 contra_subsetD in_set_conv_decomp_first self_append_conv2 set_append)
26073
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1053
  next
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1054
    assume "\<not> P x"
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1055
    hence "\<exists>x\<in>set xs. P x" using Cons(2) by simp
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1056
    thus ?thesis using `\<not> P x` Cons(1) by (metis append_Cons set_ConsD)
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1057
  qed
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1058
qed
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1059
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1060
lemma split_list_first_propE:
26734
a92057c1ee21 dropped some metis calls
haftmann
parents: 26584
diff changeset
  1061
  assumes "\<exists>x \<in> set xs. P x"
a92057c1ee21 dropped some metis calls
haftmann
parents: 26584
diff changeset
  1062
  obtains ys x zs where "xs = ys @ x # zs" and "P x" and "\<forall>y \<in> set ys. \<not> P y"
a92057c1ee21 dropped some metis calls
haftmann
parents: 26584
diff changeset
  1063
using split_list_first_prop [OF assms] by blast
26073
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1064
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1065
lemma split_list_first_prop_iff:
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1066
  "(\<exists>x \<in> set xs. P x) \<longleftrightarrow>
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1067
   (\<exists>ys x zs. xs = ys@x#zs \<and> P x \<and> (\<forall>y \<in> set ys. \<not> P y))"
26734
a92057c1ee21 dropped some metis calls
haftmann
parents: 26584
diff changeset
  1068
by (rule, erule split_list_first_prop) auto
26073
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1069
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1070
lemma split_list_last_prop:
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1071
  "\<exists>x \<in> set xs. P x \<Longrightarrow>
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1072
   \<exists>ys x zs. xs = ys@x#zs \<and> P x \<and> (\<forall>z \<in> set zs. \<not> P z)"
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1073
proof(induct xs rule:rev_induct)
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1074
  case Nil thus ?case by simp
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1075
next
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1076
  case (snoc x xs)
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1077
  show ?case
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1078
  proof cases
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1079
    assume "P x" thus ?thesis by (metis emptyE set_empty)
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1080
  next
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1081
    assume "\<not> P x"
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1082
    hence "\<exists>x\<in>set xs. P x" using snoc(2) by simp
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44635
diff changeset
  1083
    thus ?thesis using `\<not> P x` snoc(1) by fastforce
26073
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1084
  qed
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1085
qed
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1086
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1087
lemma split_list_last_propE:
26734
a92057c1ee21 dropped some metis calls
haftmann
parents: 26584
diff changeset
  1088
  assumes "\<exists>x \<in> set xs. P x"
a92057c1ee21 dropped some metis calls
haftmann
parents: 26584
diff changeset
  1089
  obtains ys x zs where "xs = ys @ x # zs" and "P x" and "\<forall>z \<in> set zs. \<not> P z"
a92057c1ee21 dropped some metis calls
haftmann
parents: 26584
diff changeset
  1090
using split_list_last_prop [OF assms] by blast
26073
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1091
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1092
lemma split_list_last_prop_iff:
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1093
  "(\<exists>x \<in> set xs. P x) \<longleftrightarrow>
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1094
   (\<exists>ys x zs. xs = ys@x#zs \<and> P x \<and> (\<forall>z \<in> set zs. \<not> P z))"
26734
a92057c1ee21 dropped some metis calls
haftmann
parents: 26584
diff changeset
  1095
by (metis split_list_last_prop [where P=P] in_set_conv_decomp)
26073
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1096
0e70d3bd2eb4 more lemmas
nipkow
parents: 25966
diff changeset
  1097
lemma finite_list: "finite A ==> EX xs. set xs = A"
26734
a92057c1ee21 dropped some metis calls
haftmann
parents: 26584
diff changeset
  1098
  by (erule finite_induct)
a92057c1ee21 dropped some metis calls
haftmann
parents: 26584
diff changeset
  1099
    (auto simp add: set.simps(2) [symmetric] simp del: set.simps(2))
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents: 13480
diff changeset
  1100
14388
04f04408b99b lemmas about card (set xs)
kleing
parents: 14343
diff changeset
  1101
lemma card_length: "card (set xs) \<le> length xs"
04f04408b99b lemmas about card (set xs)
kleing
parents: 14343
diff changeset
  1102
by (induct xs) (auto simp add: card_insert_if)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1103
26442
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1104
lemma set_minus_filter_out:
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1105
  "set xs - {y} = set (filter (\<lambda>x. \<not> (x = y)) xs)"
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1106
  by (induct xs) auto
15168
33a08cfc3ae5 new functions for sets of lists
paulson
parents: 15140
diff changeset
  1107
35115
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
  1108
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15307
diff changeset
  1109
subsubsection {* @{text filter} *}
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1110
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  1111
lemma filter_append [simp]: "filter P (xs @ ys) = filter P xs @ filter P ys"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1112
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1113
15305
0bd9eedaa301 added lemmas
nipkow
parents: 15304
diff changeset
  1114
lemma rev_filter: "rev (filter P xs) = filter P (rev xs)"
0bd9eedaa301 added lemmas
nipkow
parents: 15304
diff changeset
  1115
by (induct xs) simp_all
0bd9eedaa301 added lemmas
nipkow
parents: 15304
diff changeset
  1116
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  1117
lemma filter_filter [simp]: "filter P (filter Q xs) = filter (\<lambda>x. Q x \<and> P x) xs"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1118
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1119
16998
e0050191e2d1 Added filter lemma
nipkow
parents: 16973
diff changeset
  1120
lemma length_filter_le [simp]: "length (filter P xs) \<le> length xs"
e0050191e2d1 Added filter lemma
nipkow
parents: 16973
diff changeset
  1121
by (induct xs) (auto simp add: le_SucI)
e0050191e2d1 Added filter lemma
nipkow
parents: 16973
diff changeset
  1122
18423
d7859164447f new lemmas
nipkow
parents: 18336
diff changeset
  1123
lemma sum_length_filter_compl:
d7859164447f new lemmas
nipkow
parents: 18336
diff changeset
  1124
  "length(filter P xs) + length(filter (%x. ~P x) xs) = length xs"
d7859164447f new lemmas
nipkow
parents: 18336
diff changeset
  1125
by(induct xs) simp_all
d7859164447f new lemmas
nipkow
parents: 18336
diff changeset
  1126
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  1127
lemma filter_True [simp]: "\<forall>x \<in> set xs. P x ==> filter P xs = xs"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1128
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1129
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  1130
lemma filter_False [simp]: "\<forall>x \<in> set xs. \<not> P x ==> filter P xs = []"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1131
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1132
16998
e0050191e2d1 Added filter lemma
nipkow
parents: 16973
diff changeset
  1133
lemma filter_empty_conv: "(filter P xs = []) = (\<forall>x\<in>set xs. \<not> P x)" 
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  1134
by (induct xs) simp_all
16998
e0050191e2d1 Added filter lemma
nipkow
parents: 16973
diff changeset
  1135
e0050191e2d1 Added filter lemma
nipkow
parents: 16973
diff changeset
  1136
lemma filter_id_conv: "(filter P xs = xs) = (\<forall>x\<in>set xs. P x)"
e0050191e2d1 Added filter lemma
nipkow
parents: 16973
diff changeset
  1137
apply (induct xs)
e0050191e2d1 Added filter lemma
nipkow
parents: 16973
diff changeset
  1138
 apply auto
e0050191e2d1 Added filter lemma
nipkow
parents: 16973
diff changeset
  1139
apply(cut_tac P=P and xs=xs in length_filter_le)
e0050191e2d1 Added filter lemma
nipkow
parents: 16973
diff changeset
  1140
apply simp
e0050191e2d1 Added filter lemma
nipkow
parents: 16973
diff changeset
  1141
done
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1142
16965
46697fa536ab added map_filter lemmas
nipkow
parents: 16770
diff changeset
  1143
lemma filter_map:
46697fa536ab added map_filter lemmas
nipkow
parents: 16770
diff changeset
  1144
  "filter P (map f xs) = map f (filter (P o f) xs)"
46697fa536ab added map_filter lemmas
nipkow
parents: 16770
diff changeset
  1145
by (induct xs) simp_all
46697fa536ab added map_filter lemmas
nipkow
parents: 16770
diff changeset
  1146
46697fa536ab added map_filter lemmas
nipkow
parents: 16770
diff changeset
  1147
lemma length_filter_map[simp]:
46697fa536ab added map_filter lemmas
nipkow
parents: 16770
diff changeset
  1148
  "length (filter P (map f xs)) = length(filter (P o f) xs)"
46697fa536ab added map_filter lemmas
nipkow
parents: 16770
diff changeset
  1149
by (simp add:filter_map)
46697fa536ab added map_filter lemmas
nipkow
parents: 16770
diff changeset
  1150
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  1151
lemma filter_is_subset [simp]: "set (filter P xs) \<le> set xs"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1152
by auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1153
15246
0984a2c2868b added and renamed
nipkow
parents: 15245
diff changeset
  1154
lemma length_filter_less:
0984a2c2868b added and renamed
nipkow
parents: 15245
diff changeset
  1155
  "\<lbrakk> x : set xs; ~ P x \<rbrakk> \<Longrightarrow> length(filter P xs) < length xs"
0984a2c2868b added and renamed
nipkow
parents: 15245
diff changeset
  1156
proof (induct xs)
0984a2c2868b added and renamed
nipkow
parents: 15245
diff changeset
  1157
  case Nil thus ?case by simp
0984a2c2868b added and renamed
nipkow
parents: 15245
diff changeset
  1158
next
0984a2c2868b added and renamed
nipkow
parents: 15245
diff changeset
  1159
  case (Cons x xs) thus ?case
0984a2c2868b added and renamed
nipkow
parents: 15245
diff changeset
  1160
    apply (auto split:split_if_asm)
0984a2c2868b added and renamed
nipkow
parents: 15245
diff changeset
  1161
    using length_filter_le[of P xs] apply arith
0984a2c2868b added and renamed
nipkow
parents: 15245
diff changeset
  1162
  done
0984a2c2868b added and renamed
nipkow
parents: 15245
diff changeset
  1163
qed
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1164
15281
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  1165
lemma length_filter_conv_card:
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  1166
 "length(filter p xs) = card{i. i < length xs & p(xs!i)}"
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  1167
proof (induct xs)
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  1168
  case Nil thus ?case by simp
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  1169
next
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  1170
  case (Cons x xs)
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  1171
  let ?S = "{i. i < length xs & p(xs!i)}"
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  1172
  have fin: "finite ?S" by(fast intro: bounded_nat_set_is_finite)
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  1173
  show ?case (is "?l = card ?S'")
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  1174
  proof (cases)
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  1175
    assume "p x"
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  1176
    hence eq: "?S' = insert 0 (Suc ` ?S)"
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25157
diff changeset
  1177
      by(auto simp: image_def split:nat.split dest:gr0_implies_Suc)
15281
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  1178
    have "length (filter p (x # xs)) = Suc(card ?S)"
23388
77645da0db85 tuned proofs: avoid implicit prems;
wenzelm
parents: 23279
diff changeset
  1179
      using Cons `p x` by simp
15281
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  1180
    also have "\<dots> = Suc(card(Suc ` ?S))" using fin
44921
58eef4843641 tuned proofs
huffman
parents: 44916
diff changeset
  1181
      by (simp add: card_image)
15281
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  1182
    also have "\<dots> = card ?S'" using eq fin
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  1183
      by (simp add:card_insert_if) (simp add:image_def)
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  1184
    finally show ?thesis .
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  1185
  next
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  1186
    assume "\<not> p x"
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  1187
    hence eq: "?S' = Suc ` ?S"
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25157
diff changeset
  1188
      by(auto simp add: image_def split:nat.split elim:lessE)
15281
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  1189
    have "length (filter p (x # xs)) = card ?S"
23388
77645da0db85 tuned proofs: avoid implicit prems;
wenzelm
parents: 23279
diff changeset
  1190
      using Cons `\<not> p x` by simp
15281
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  1191
    also have "\<dots> = card(Suc ` ?S)" using fin
44921
58eef4843641 tuned proofs
huffman
parents: 44916
diff changeset
  1192
      by (simp add: card_image)
15281
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  1193
    also have "\<dots> = card ?S'" using eq fin
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  1194
      by (simp add:card_insert_if)
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  1195
    finally show ?thesis .
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  1196
  qed
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  1197
qed
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  1198
17629
f8ea8068c6d9 a few new filter lemmas
nipkow
parents: 17589
diff changeset
  1199
lemma Cons_eq_filterD:
f8ea8068c6d9 a few new filter lemmas
nipkow
parents: 17589
diff changeset
  1200
 "x#xs = filter P ys \<Longrightarrow>
f8ea8068c6d9 a few new filter lemmas
nipkow
parents: 17589
diff changeset
  1201
  \<exists>us vs. ys = us @ x # vs \<and> (\<forall>u\<in>set us. \<not> P u) \<and> P x \<and> xs = filter P vs"
19585
70a1ce3b23ae removed 'concl is' patterns;
wenzelm
parents: 19487
diff changeset
  1202
  (is "_ \<Longrightarrow> \<exists>us vs. ?P ys us vs")
17629
f8ea8068c6d9 a few new filter lemmas
nipkow
parents: 17589
diff changeset
  1203
proof(induct ys)
f8ea8068c6d9 a few new filter lemmas
nipkow
parents: 17589
diff changeset
  1204
  case Nil thus ?case by simp
f8ea8068c6d9 a few new filter lemmas
nipkow
parents: 17589
diff changeset
  1205
next
f8ea8068c6d9 a few new filter lemmas
nipkow
parents: 17589
diff changeset
  1206
  case (Cons y ys)
f8ea8068c6d9 a few new filter lemmas
nipkow
parents: 17589
diff changeset
  1207
  show ?case (is "\<exists>x. ?Q x")
f8ea8068c6d9 a few new filter lemmas
nipkow
parents: 17589
diff changeset
  1208
  proof cases
f8ea8068c6d9 a few new filter lemmas
nipkow
parents: 17589
diff changeset
  1209
    assume Py: "P y"
f8ea8068c6d9 a few new filter lemmas
nipkow
parents: 17589
diff changeset
  1210
    show ?thesis
f8ea8068c6d9 a few new filter lemmas
nipkow
parents: 17589
diff changeset
  1211
    proof cases
25221
5ded95dda5df append/member: more light-weight way to declare authentic syntax;
wenzelm
parents: 25215
diff changeset
  1212
      assume "x = y"
5ded95dda5df append/member: more light-weight way to declare authentic syntax;
wenzelm
parents: 25215
diff changeset
  1213
      with Py Cons.prems have "?Q []" by simp
5ded95dda5df append/member: more light-weight way to declare authentic syntax;
wenzelm
parents: 25215
diff changeset
  1214
      then show ?thesis ..
17629
f8ea8068c6d9 a few new filter lemmas
nipkow
parents: 17589
diff changeset
  1215
    next
25221
5ded95dda5df append/member: more light-weight way to declare authentic syntax;
wenzelm
parents: 25215
diff changeset
  1216
      assume "x \<noteq> y"
5ded95dda5df append/member: more light-weight way to declare authentic syntax;
wenzelm
parents: 25215
diff changeset
  1217
      with Py Cons.prems show ?thesis by simp
17629
f8ea8068c6d9 a few new filter lemmas
nipkow
parents: 17589
diff changeset
  1218
    qed
f8ea8068c6d9 a few new filter lemmas
nipkow
parents: 17589
diff changeset
  1219
  next
25221
5ded95dda5df append/member: more light-weight way to declare authentic syntax;
wenzelm
parents: 25215
diff changeset
  1220
    assume "\<not> P y"
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44635
diff changeset
  1221
    with Cons obtain us vs where "?P (y#ys) (y#us) vs" by fastforce
25221
5ded95dda5df append/member: more light-weight way to declare authentic syntax;
wenzelm
parents: 25215
diff changeset
  1222
    then have "?Q (y#us)" by simp
5ded95dda5df append/member: more light-weight way to declare authentic syntax;
wenzelm
parents: 25215
diff changeset
  1223
    then show ?thesis ..
17629
f8ea8068c6d9 a few new filter lemmas
nipkow
parents: 17589
diff changeset
  1224
  qed
f8ea8068c6d9 a few new filter lemmas
nipkow
parents: 17589
diff changeset
  1225
qed
f8ea8068c6d9 a few new filter lemmas
nipkow
parents: 17589
diff changeset
  1226
f8ea8068c6d9 a few new filter lemmas
nipkow
parents: 17589
diff changeset
  1227
lemma filter_eq_ConsD:
f8ea8068c6d9 a few new filter lemmas
nipkow
parents: 17589
diff changeset
  1228
 "filter P ys = x#xs \<Longrightarrow>
f8ea8068c6d9 a few new filter lemmas
nipkow
parents: 17589
diff changeset
  1229
  \<exists>us vs. ys = us @ x # vs \<and> (\<forall>u\<in>set us. \<not> P u) \<and> P x \<and> xs = filter P vs"
f8ea8068c6d9 a few new filter lemmas
nipkow
parents: 17589
diff changeset
  1230
by(rule Cons_eq_filterD) simp
f8ea8068c6d9 a few new filter lemmas
nipkow
parents: 17589
diff changeset
  1231
f8ea8068c6d9 a few new filter lemmas
nipkow
parents: 17589
diff changeset
  1232
lemma filter_eq_Cons_iff:
f8ea8068c6d9 a few new filter lemmas
nipkow
parents: 17589
diff changeset
  1233
 "(filter P ys = x#xs) =
f8ea8068c6d9 a few new filter lemmas
nipkow
parents: 17589
diff changeset
  1234
  (\<exists>us vs. ys = us @ x # vs \<and> (\<forall>u\<in>set us. \<not> P u) \<and> P x \<and> xs = filter P vs)"
f8ea8068c6d9 a few new filter lemmas
nipkow
parents: 17589
diff changeset
  1235
by(auto dest:filter_eq_ConsD)
f8ea8068c6d9 a few new filter lemmas
nipkow
parents: 17589
diff changeset
  1236
f8ea8068c6d9 a few new filter lemmas
nipkow
parents: 17589
diff changeset
  1237
lemma Cons_eq_filter_iff:
f8ea8068c6d9 a few new filter lemmas
nipkow
parents: 17589
diff changeset
  1238
 "(x#xs = filter P ys) =
f8ea8068c6d9 a few new filter lemmas
nipkow
parents: 17589
diff changeset
  1239
  (\<exists>us vs. ys = us @ x # vs \<and> (\<forall>u\<in>set us. \<not> P u) \<and> P x \<and> xs = filter P vs)"
f8ea8068c6d9 a few new filter lemmas
nipkow
parents: 17589
diff changeset
  1240
by(auto dest:Cons_eq_filterD)
f8ea8068c6d9 a few new filter lemmas
nipkow
parents: 17589
diff changeset
  1241
44013
5cfc1c36ae97 moved recdef package to HOL/Library/Old_Recdef.thy
krauss
parents: 43594
diff changeset
  1242
lemma filter_cong[fundef_cong]:
17501
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1243
 "xs = ys \<Longrightarrow> (\<And>x. x \<in> set ys \<Longrightarrow> P x = Q x) \<Longrightarrow> filter P xs = filter Q ys"
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1244
apply simp
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1245
apply(erule thin_rl)
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1246
by (induct ys) simp_all
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1247
15281
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  1248
26442
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1249
subsubsection {* List partitioning *}
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1250
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1251
primrec partition :: "('a \<Rightarrow> bool) \<Rightarrow>'a list \<Rightarrow> 'a list \<times> 'a list" where
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1252
  "partition P [] = ([], [])"
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1253
  | "partition P (x # xs) = 
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1254
      (let (yes, no) = partition P xs
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1255
      in if P x then (x # yes, no) else (yes, x # no))"
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1256
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1257
lemma partition_filter1:
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1258
    "fst (partition P xs) = filter P xs"
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1259
by (induct xs) (auto simp add: Let_def split_def)
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1260
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1261
lemma partition_filter2:
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1262
    "snd (partition P xs) = filter (Not o P) xs"
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1263
by (induct xs) (auto simp add: Let_def split_def)
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1264
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1265
lemma partition_P:
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1266
  assumes "partition P xs = (yes, no)"
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1267
  shows "(\<forall>p \<in> set yes.  P p) \<and> (\<forall>p  \<in> set no. \<not> P p)"
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1268
proof -
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1269
  from assms have "yes = fst (partition P xs)" and "no = snd (partition P xs)"
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1270
    by simp_all
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1271
  then show ?thesis by (simp_all add: partition_filter1 partition_filter2)
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1272
qed
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1273
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1274
lemma partition_set:
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1275
  assumes "partition P xs = (yes, no)"
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1276
  shows "set yes \<union> set no = set xs"
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1277
proof -
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1278
  from assms have "yes = fst (partition P xs)" and "no = snd (partition P xs)"
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1279
    by simp_all
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1280
  then show ?thesis by (auto simp add: partition_filter1 partition_filter2) 
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1281
qed
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1282
33639
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1283
lemma partition_filter_conv[simp]:
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1284
  "partition f xs = (filter f xs,filter (Not o f) xs)"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1285
unfolding partition_filter2[symmetric]
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1286
unfolding partition_filter1[symmetric] by simp
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1287
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1288
declare partition.simps[simp del]
26442
57fb6a8b099e restructuring; explicit case names for rule list_induct2
haftmann
parents: 26300
diff changeset
  1289
35115
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
  1290
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15307
diff changeset
  1291
subsubsection {* @{text concat} *}
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1292
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  1293
lemma concat_append [simp]: "concat (xs @ ys) = concat xs @ concat ys"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1294
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1295
18447
da548623916a removed or modified some instances of [iff]
paulson
parents: 18423
diff changeset
  1296
lemma concat_eq_Nil_conv [simp]: "(concat xss = []) = (\<forall>xs \<in> set xss. xs = [])"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1297
by (induct xss) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1298
18447
da548623916a removed or modified some instances of [iff]
paulson
parents: 18423
diff changeset
  1299
lemma Nil_eq_concat_conv [simp]: "([] = concat xss) = (\<forall>xs \<in> set xss. xs = [])"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1300
by (induct xss) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1301
24308
700e745994c1 removed set_concat_map and improved set_concat
nipkow
parents: 24286
diff changeset
  1302
lemma set_concat [simp]: "set (concat xs) = (UN x:set xs. set x)"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1303
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1304
24476
f7ad9fbbeeaa turned list comprehension translations into ML to optimize base case
nipkow
parents: 24471
diff changeset
  1305
lemma concat_map_singleton[simp]: "concat(map (%x. [f x]) xs) = map f xs"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  1306
by (induct xs) auto
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  1307
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  1308
lemma map_concat: "map f (concat xs) = concat (map (map f) xs)"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1309
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1310
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  1311
lemma filter_concat: "filter p (concat xs) = concat (map (filter p) xs)"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1312
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1313
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  1314
lemma rev_concat: "rev (concat xs) = concat (map rev (rev xs))"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1315
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1316
40365
a1456f2e1c9d added two lemmas about injectivity of concat to the list theory
bulwahn
parents: 40304
diff changeset
  1317
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)"
a1456f2e1c9d added two lemmas about injectivity of concat to the list theory
bulwahn
parents: 40304
diff changeset
  1318
proof (induct xs arbitrary: ys)
a1456f2e1c9d added two lemmas about injectivity of concat to the list theory
bulwahn
parents: 40304
diff changeset
  1319
  case (Cons x xs ys)
a1456f2e1c9d added two lemmas about injectivity of concat to the list theory
bulwahn
parents: 40304
diff changeset
  1320
  thus ?case by (cases ys) auto
a1456f2e1c9d added two lemmas about injectivity of concat to the list theory
bulwahn
parents: 40304
diff changeset
  1321
qed (auto)
a1456f2e1c9d added two lemmas about injectivity of concat to the list theory
bulwahn
parents: 40304
diff changeset
  1322
a1456f2e1c9d added two lemmas about injectivity of concat to the list theory
bulwahn
parents: 40304
diff changeset
  1323
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"
a1456f2e1c9d added two lemmas about injectivity of concat to the list theory
bulwahn
parents: 40304
diff changeset
  1324
by (simp add: concat_eq_concat_iff)
a1456f2e1c9d added two lemmas about injectivity of concat to the list theory
bulwahn
parents: 40304
diff changeset
  1325
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1326
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15307
diff changeset
  1327
subsubsection {* @{text nth} *}
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1328
29827
c82b3e8a4daf code theorems for nth, list_update
haftmann
parents: 29822
diff changeset
  1329
lemma nth_Cons_0 [simp, code]: "(x # xs)!0 = x"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1330
by auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1331
29827
c82b3e8a4daf code theorems for nth, list_update
haftmann
parents: 29822
diff changeset
  1332
lemma nth_Cons_Suc [simp, code]: "(x # xs)!(Suc n) = xs!n"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1333
by auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1334
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  1335
declare nth.simps [simp del]
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1336
41842
d8f76db6a207 added simp lemma nth_Cons_pos to List
nipkow
parents: 41697
diff changeset
  1337
lemma nth_Cons_pos[simp]: "0 < n \<Longrightarrow> (x#xs) ! n = xs ! (n - 1)"
d8f76db6a207 added simp lemma nth_Cons_pos to List
nipkow
parents: 41697
diff changeset
  1338
by(auto simp: Nat.gr0_conv_Suc)
d8f76db6a207 added simp lemma nth_Cons_pos to List
nipkow
parents: 41697
diff changeset
  1339
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1340
lemma nth_append:
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1341
  "(xs @ ys)!n = (if n < length xs then xs!n else ys!(n - length xs))"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1342
apply (induct xs arbitrary: n, simp)
14208
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
  1343
apply (case_tac n, auto)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1344
done
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1345
14402
4201e1916482 moved lemmas from MicroJava/Comp/AuxLemmas.thy to List.thy
nipkow
parents: 14395
diff changeset
  1346
lemma nth_append_length [simp]: "(xs @ x # ys) ! length xs = x"
25221
5ded95dda5df append/member: more light-weight way to declare authentic syntax;
wenzelm
parents: 25215
diff changeset
  1347
by (induct xs) auto
14402
4201e1916482 moved lemmas from MicroJava/Comp/AuxLemmas.thy to List.thy
nipkow
parents: 14395
diff changeset
  1348
4201e1916482 moved lemmas from MicroJava/Comp/AuxLemmas.thy to List.thy
nipkow
parents: 14395
diff changeset
  1349
lemma nth_append_length_plus[simp]: "(xs @ ys) ! (length xs + n) = ys ! n"
25221
5ded95dda5df append/member: more light-weight way to declare authentic syntax;
wenzelm
parents: 25215
diff changeset
  1350
by (induct xs) auto
14402
4201e1916482 moved lemmas from MicroJava/Comp/AuxLemmas.thy to List.thy
nipkow
parents: 14395
diff changeset
  1351
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1352
lemma nth_map [simp]: "n < length xs ==> (map f xs)!n = f(xs!n)"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1353
apply (induct xs arbitrary: n, simp)
14208
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
  1354
apply (case_tac n, auto)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1355
done
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1356
18423
d7859164447f new lemmas
nipkow
parents: 18336
diff changeset
  1357
lemma hd_conv_nth: "xs \<noteq> [] \<Longrightarrow> hd xs = xs!0"
d7859164447f new lemmas
nipkow
parents: 18336
diff changeset
  1358
by(cases xs) simp_all
d7859164447f new lemmas
nipkow
parents: 18336
diff changeset
  1359
18049
156bba334c12 A few new lemmas
nipkow
parents: 17956
diff changeset
  1360
156bba334c12 A few new lemmas
nipkow
parents: 17956
diff changeset
  1361
lemma list_eq_iff_nth_eq:
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1362
 "(xs = ys) = (length xs = length ys \<and> (ALL i<length xs. xs!i = ys!i))"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1363
apply(induct xs arbitrary: ys)
24632
779fc4fcbf8b metis now available in PreList
paulson
parents: 24617
diff changeset
  1364
 apply force
18049
156bba334c12 A few new lemmas
nipkow
parents: 17956
diff changeset
  1365
apply(case_tac ys)
156bba334c12 A few new lemmas
nipkow
parents: 17956
diff changeset
  1366
 apply simp
156bba334c12 A few new lemmas
nipkow
parents: 17956
diff changeset
  1367
apply(simp add:nth_Cons split:nat.split)apply blast
156bba334c12 A few new lemmas
nipkow
parents: 17956
diff changeset
  1368
done
156bba334c12 A few new lemmas
nipkow
parents: 17956
diff changeset
  1369
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  1370
lemma set_conv_nth: "set xs = {xs!i | i. i < length xs}"
15251
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15246
diff changeset
  1371
apply (induct xs, simp, simp)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1372
apply safe
24632
779fc4fcbf8b metis now available in PreList
paulson
parents: 24617
diff changeset
  1373
apply (metis nat_case_0 nth.simps zero_less_Suc)
779fc4fcbf8b metis now available in PreList
paulson
parents: 24617
diff changeset
  1374
apply (metis less_Suc_eq_0_disj nth_Cons_Suc)
14208
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
  1375
apply (case_tac i, simp)
24632
779fc4fcbf8b metis now available in PreList
paulson
parents: 24617
diff changeset
  1376
apply (metis diff_Suc_Suc nat_case_Suc nth.simps zero_less_diff)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1377
done
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1378
17501
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1379
lemma in_set_conv_nth: "(x \<in> set xs) = (\<exists>i < length xs. xs!i = x)"
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1380
by(auto simp:set_conv_nth)
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1381
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1382
lemma list_ball_nth: "[| n < length xs; !x : set xs. P x|] ==> P(xs!n)"
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1383
by (auto simp add: set_conv_nth)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1384
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  1385
lemma nth_mem [simp]: "n < length xs ==> xs!n : set xs"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1386
by (auto simp add: set_conv_nth)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1387
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1388
lemma all_nth_imp_all_set:
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1389
"[| !i < length xs. P(xs!i); x : set xs|] ==> P x"
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1390
by (auto simp add: set_conv_nth)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1391
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1392
lemma all_set_conv_all_nth:
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1393
"(\<forall>x \<in> set xs. P x) = (\<forall>i. i < length xs --> P (xs ! i))"
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1394
by (auto simp add: set_conv_nth)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1395
25296
c187b7422156 rev_nth
kleing
parents: 25287
diff changeset
  1396
lemma rev_nth:
c187b7422156 rev_nth
kleing
parents: 25287
diff changeset
  1397
  "n < size xs \<Longrightarrow> rev xs ! n = xs ! (length xs - Suc n)"
c187b7422156 rev_nth
kleing
parents: 25287
diff changeset
  1398
proof (induct xs arbitrary: n)
c187b7422156 rev_nth
kleing
parents: 25287
diff changeset
  1399
  case Nil thus ?case by simp
c187b7422156 rev_nth
kleing
parents: 25287
diff changeset
  1400
next
c187b7422156 rev_nth
kleing
parents: 25287
diff changeset
  1401
  case (Cons x xs)
c187b7422156 rev_nth
kleing
parents: 25287
diff changeset
  1402
  hence n: "n < Suc (length xs)" by simp
c187b7422156 rev_nth
kleing
parents: 25287
diff changeset
  1403
  moreover
c187b7422156 rev_nth
kleing
parents: 25287
diff changeset
  1404
  { assume "n < length xs"
c187b7422156 rev_nth
kleing
parents: 25287
diff changeset
  1405
    with n obtain n' where "length xs - n = Suc n'"
c187b7422156 rev_nth
kleing
parents: 25287
diff changeset
  1406
      by (cases "length xs - n", auto)
c187b7422156 rev_nth
kleing
parents: 25287
diff changeset
  1407
    moreover
c187b7422156 rev_nth
kleing
parents: 25287
diff changeset
  1408
    then have "length xs - Suc n = n'" by simp
c187b7422156 rev_nth
kleing
parents: 25287
diff changeset
  1409
    ultimately
c187b7422156 rev_nth
kleing
parents: 25287
diff changeset
  1410
    have "xs ! (length xs - Suc n) = (x # xs) ! (length xs - n)" by simp
c187b7422156 rev_nth
kleing
parents: 25287
diff changeset
  1411
  }
c187b7422156 rev_nth
kleing
parents: 25287
diff changeset
  1412
  ultimately
c187b7422156 rev_nth
kleing
parents: 25287
diff changeset
  1413
  show ?case by (clarsimp simp add: Cons nth_append)
c187b7422156 rev_nth
kleing
parents: 25287
diff changeset
  1414
qed
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1415
31159
bac0d673b6d6 new lemma
nipkow
parents: 31154
diff changeset
  1416
lemma Skolem_list_nth:
bac0d673b6d6 new lemma
nipkow
parents: 31154
diff changeset
  1417
  "(ALL i<k. EX x. P i x) = (EX xs. size xs = k & (ALL i<k. P i (xs!i)))"
bac0d673b6d6 new lemma
nipkow
parents: 31154
diff changeset
  1418
  (is "_ = (EX xs. ?P k xs)")
bac0d673b6d6 new lemma
nipkow
parents: 31154
diff changeset
  1419
proof(induct k)
bac0d673b6d6 new lemma
nipkow
parents: 31154
diff changeset
  1420
  case 0 show ?case by simp
bac0d673b6d6 new lemma
nipkow
parents: 31154
diff changeset
  1421
next
bac0d673b6d6 new lemma
nipkow
parents: 31154
diff changeset
  1422
  case (Suc k)
bac0d673b6d6 new lemma
nipkow
parents: 31154
diff changeset
  1423
  show ?case (is "?L = ?R" is "_ = (EX xs. ?P' xs)")
bac0d673b6d6 new lemma
nipkow
parents: 31154
diff changeset
  1424
  proof
bac0d673b6d6 new lemma
nipkow
parents: 31154
diff changeset
  1425
    assume "?R" thus "?L" using Suc by auto
bac0d673b6d6 new lemma
nipkow
parents: 31154
diff changeset
  1426
  next
bac0d673b6d6 new lemma
nipkow
parents: 31154
diff changeset
  1427
    assume "?L"
bac0d673b6d6 new lemma
nipkow
parents: 31154
diff changeset
  1428
    with Suc obtain x xs where "?P k xs & P k x" by (metis less_Suc_eq)
bac0d673b6d6 new lemma
nipkow
parents: 31154
diff changeset
  1429
    hence "?P'(xs@[x])" by(simp add:nth_append less_Suc_eq)
bac0d673b6d6 new lemma
nipkow
parents: 31154
diff changeset
  1430
    thus "?R" ..
bac0d673b6d6 new lemma
nipkow
parents: 31154
diff changeset
  1431
  qed
bac0d673b6d6 new lemma
nipkow
parents: 31154
diff changeset
  1432
qed
bac0d673b6d6 new lemma
nipkow
parents: 31154
diff changeset
  1433
bac0d673b6d6 new lemma
nipkow
parents: 31154
diff changeset
  1434
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15307
diff changeset
  1435
subsubsection {* @{text list_update} *}
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1436
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1437
lemma length_list_update [simp]: "length(xs[i:=x]) = length xs"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1438
by (induct xs arbitrary: i) (auto split: nat.split)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1439
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1440
lemma nth_list_update:
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1441
"i < length xs==> (xs[i:=x])!j = (if i = j then x else xs!j)"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1442
by (induct xs arbitrary: i j) (auto simp add: nth_Cons split: nat.split)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1443
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  1444
lemma nth_list_update_eq [simp]: "i < length xs ==> (xs[i:=x])!i = x"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1445
by (simp add: nth_list_update)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1446
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1447
lemma nth_list_update_neq [simp]: "i \<noteq> j ==> xs[i:=x]!j = xs!j"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1448
by (induct xs arbitrary: i j) (auto simp add: nth_Cons split: nat.split)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1449
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1450
lemma list_update_id[simp]: "xs[i := xs!i] = xs"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1451
by (induct xs arbitrary: i) (simp_all split:nat.splits)
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1452
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1453
lemma list_update_beyond[simp]: "length xs \<le> i \<Longrightarrow> xs[i:=x] = xs"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1454
apply (induct xs arbitrary: i)
17501
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1455
 apply simp
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1456
apply (case_tac i)
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1457
apply simp_all
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1458
done
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1459
31077
28dd6fd3d184 more lemmas
nipkow
parents: 31055
diff changeset
  1460
lemma list_update_nonempty[simp]: "xs[k:=x] = [] \<longleftrightarrow> xs=[]"
28dd6fd3d184 more lemmas
nipkow
parents: 31055
diff changeset
  1461
by(metis length_0_conv length_list_update)
28dd6fd3d184 more lemmas
nipkow
parents: 31055
diff changeset
  1462
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1463
lemma list_update_same_conv:
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1464
"i < length xs ==> (xs[i := x] = xs) = (xs!i = x)"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1465
by (induct xs arbitrary: i) (auto split: nat.split)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1466
14187
26dfcd0ac436 Added new theorems
nipkow
parents: 14111
diff changeset
  1467
lemma list_update_append1:
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1468
 "i < size xs \<Longrightarrow> (xs @ ys)[i:=x] = xs[i:=x] @ ys"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1469
apply (induct xs arbitrary: i, simp)
14187
26dfcd0ac436 Added new theorems
nipkow
parents: 14111
diff changeset
  1470
apply(simp split:nat.split)
26dfcd0ac436 Added new theorems
nipkow
parents: 14111
diff changeset
  1471
done
26dfcd0ac436 Added new theorems
nipkow
parents: 14111
diff changeset
  1472
15868
9634b3f9d910 more about list_update
kleing
parents: 15693
diff changeset
  1473
lemma list_update_append:
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1474
  "(xs @ ys) [n:= x] = 
15868
9634b3f9d910 more about list_update
kleing
parents: 15693
diff changeset
  1475
  (if n < length xs then xs[n:= x] @ ys else xs @ (ys [n-length xs:= x]))"
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1476
by (induct xs arbitrary: n) (auto split:nat.splits)
15868
9634b3f9d910 more about list_update
kleing
parents: 15693
diff changeset
  1477
14402
4201e1916482 moved lemmas from MicroJava/Comp/AuxLemmas.thy to List.thy
nipkow
parents: 14395
diff changeset
  1478
lemma list_update_length [simp]:
4201e1916482 moved lemmas from MicroJava/Comp/AuxLemmas.thy to List.thy
nipkow
parents: 14395
diff changeset
  1479
 "(xs @ x # ys)[length xs := y] = (xs @ y # ys)"
4201e1916482 moved lemmas from MicroJava/Comp/AuxLemmas.thy to List.thy
nipkow
parents: 14395
diff changeset
  1480
by (induct xs, auto)
4201e1916482 moved lemmas from MicroJava/Comp/AuxLemmas.thy to List.thy
nipkow
parents: 14395
diff changeset
  1481
31264
2662d1cdc51f more lemmas
nipkow
parents: 31258
diff changeset
  1482
lemma map_update: "map f (xs[k:= y]) = (map f xs)[k := f y]"
2662d1cdc51f more lemmas
nipkow
parents: 31258
diff changeset
  1483
by(induct xs arbitrary: k)(auto split:nat.splits)
2662d1cdc51f more lemmas
nipkow
parents: 31258
diff changeset
  1484
2662d1cdc51f more lemmas
nipkow
parents: 31258
diff changeset
  1485
lemma rev_update:
2662d1cdc51f more lemmas
nipkow
parents: 31258
diff changeset
  1486
  "k < length xs \<Longrightarrow> rev (xs[k:= y]) = (rev xs)[length xs - k - 1 := y]"
2662d1cdc51f more lemmas
nipkow
parents: 31258
diff changeset
  1487
by (induct xs arbitrary: k) (auto simp: list_update_append split:nat.splits)
2662d1cdc51f more lemmas
nipkow
parents: 31258
diff changeset
  1488
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1489
lemma update_zip:
31080
21ffc770ebc0 lemmas by Andreas Lochbihler
nipkow
parents: 31077
diff changeset
  1490
  "(zip xs ys)[i:=xy] = zip (xs[i:=fst xy]) (ys[i:=snd xy])"
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1491
by (induct ys arbitrary: i xy xs) (auto, case_tac xs, auto split: nat.split)
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1492
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1493
lemma set_update_subset_insert: "set(xs[i:=x]) <= insert x (set xs)"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1494
by (induct xs arbitrary: i) (auto split: nat.split)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1495
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1496
lemma set_update_subsetI: "[| set xs <= A; x:A |] ==> set(xs[i := x]) <= A"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1497
by (blast dest!: set_update_subset_insert [THEN subsetD])
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1498
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1499
lemma set_update_memI: "n < length xs \<Longrightarrow> x \<in> set (xs[n := x])"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1500
by (induct xs arbitrary: n) (auto split:nat.splits)
15868
9634b3f9d910 more about list_update
kleing
parents: 15693
diff changeset
  1501
31077
28dd6fd3d184 more lemmas
nipkow
parents: 31055
diff changeset
  1502
lemma list_update_overwrite[simp]:
24796
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  1503
  "xs [i := x, i := y] = xs [i := y]"
31077
28dd6fd3d184 more lemmas
nipkow
parents: 31055
diff changeset
  1504
apply (induct xs arbitrary: i) apply simp
28dd6fd3d184 more lemmas
nipkow
parents: 31055
diff changeset
  1505
apply (case_tac i, simp_all)
24796
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  1506
done
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  1507
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  1508
lemma list_update_swap:
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  1509
  "i \<noteq> i' \<Longrightarrow> xs [i := x, i' := x'] = xs [i' := x', i := x]"
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  1510
apply (induct xs arbitrary: i i')
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  1511
apply simp
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  1512
apply (case_tac i, case_tac i')
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  1513
apply auto
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  1514
apply (case_tac i')
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  1515
apply auto
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  1516
done
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  1517
29827
c82b3e8a4daf code theorems for nth, list_update
haftmann
parents: 29822
diff changeset
  1518
lemma list_update_code [code]:
c82b3e8a4daf code theorems for nth, list_update
haftmann
parents: 29822
diff changeset
  1519
  "[][i := y] = []"
c82b3e8a4daf code theorems for nth, list_update
haftmann
parents: 29822
diff changeset
  1520
  "(x # xs)[0 := y] = y # xs"
c82b3e8a4daf code theorems for nth, list_update
haftmann
parents: 29822
diff changeset
  1521
  "(x # xs)[Suc i := y] = x # xs[i := y]"
c82b3e8a4daf code theorems for nth, list_update
haftmann
parents: 29822
diff changeset
  1522
  by simp_all
c82b3e8a4daf code theorems for nth, list_update
haftmann
parents: 29822
diff changeset
  1523
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1524
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15307
diff changeset
  1525
subsubsection {* @{text last} and @{text butlast} *}
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1526
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  1527
lemma last_snoc [simp]: "last (xs @ [x]) = x"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1528
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1529
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  1530
lemma butlast_snoc [simp]: "butlast (xs @ [x]) = xs"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1531
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1532
14302
6c24235e8d5d *** empty log message ***
nipkow
parents: 14300
diff changeset
  1533
lemma last_ConsL: "xs = [] \<Longrightarrow> last(x#xs) = x"
44921
58eef4843641 tuned proofs
huffman
parents: 44916
diff changeset
  1534
  by simp
14302
6c24235e8d5d *** empty log message ***
nipkow
parents: 14300
diff changeset
  1535
6c24235e8d5d *** empty log message ***
nipkow
parents: 14300
diff changeset
  1536
lemma last_ConsR: "xs \<noteq> [] \<Longrightarrow> last(x#xs) = last xs"
44921
58eef4843641 tuned proofs
huffman
parents: 44916
diff changeset
  1537
  by simp
14302
6c24235e8d5d *** empty log message ***
nipkow
parents: 14300
diff changeset
  1538
6c24235e8d5d *** empty log message ***
nipkow
parents: 14300
diff changeset
  1539
lemma last_append: "last(xs @ ys) = (if ys = [] then last xs else last ys)"
6c24235e8d5d *** empty log message ***
nipkow
parents: 14300
diff changeset
  1540
by (induct xs) (auto)
6c24235e8d5d *** empty log message ***
nipkow
parents: 14300
diff changeset
  1541
6c24235e8d5d *** empty log message ***
nipkow
parents: 14300
diff changeset
  1542
lemma last_appendL[simp]: "ys = [] \<Longrightarrow> last(xs @ ys) = last xs"
6c24235e8d5d *** empty log message ***
nipkow
parents: 14300
diff changeset
  1543
by(simp add:last_append)
6c24235e8d5d *** empty log message ***
nipkow
parents: 14300
diff changeset
  1544
6c24235e8d5d *** empty log message ***
nipkow
parents: 14300
diff changeset
  1545
lemma last_appendR[simp]: "ys \<noteq> [] \<Longrightarrow> last(xs @ ys) = last ys"
6c24235e8d5d *** empty log message ***
nipkow
parents: 14300
diff changeset
  1546
by(simp add:last_append)
6c24235e8d5d *** empty log message ***
nipkow
parents: 14300
diff changeset
  1547
17762
478869f444ca new hd/rev/last lemmas
nipkow
parents: 17724
diff changeset
  1548
lemma hd_rev: "xs \<noteq> [] \<Longrightarrow> hd(rev xs) = last xs"
478869f444ca new hd/rev/last lemmas
nipkow
parents: 17724
diff changeset
  1549
by(rule rev_exhaust[of xs]) simp_all
478869f444ca new hd/rev/last lemmas
nipkow
parents: 17724
diff changeset
  1550
478869f444ca new hd/rev/last lemmas
nipkow
parents: 17724
diff changeset
  1551
lemma last_rev: "xs \<noteq> [] \<Longrightarrow> last(rev xs) = hd xs"
478869f444ca new hd/rev/last lemmas
nipkow
parents: 17724
diff changeset
  1552
by(cases xs) simp_all
478869f444ca new hd/rev/last lemmas
nipkow
parents: 17724
diff changeset
  1553
17765
e3cd31bc2e40 added last in set lemma
nipkow
parents: 17762
diff changeset
  1554
lemma last_in_set[simp]: "as \<noteq> [] \<Longrightarrow> last as \<in> set as"
e3cd31bc2e40 added last in set lemma
nipkow
parents: 17762
diff changeset
  1555
by (induct as) auto
17762
478869f444ca new hd/rev/last lemmas
nipkow
parents: 17724
diff changeset
  1556
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  1557
lemma length_butlast [simp]: "length (butlast xs) = length xs - 1"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1558
by (induct xs rule: rev_induct) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1559
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1560
lemma butlast_append:
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1561
  "butlast (xs @ ys) = (if ys = [] then butlast xs else xs @ butlast ys)"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1562
by (induct xs arbitrary: ys) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1563
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  1564
lemma append_butlast_last_id [simp]:
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1565
"xs \<noteq> [] ==> butlast xs @ [last xs] = xs"
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1566
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1567
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  1568
lemma in_set_butlastD: "x : set (butlast xs) ==> x : set xs"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1569
by (induct xs) (auto split: split_if_asm)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1570
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1571
lemma in_set_butlast_appendI:
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1572
"x : set (butlast xs) | x : set (butlast ys) ==> x : set (butlast (xs @ ys))"
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1573
by (auto dest: in_set_butlastD simp add: butlast_append)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1574
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1575
lemma last_drop[simp]: "n < length xs \<Longrightarrow> last (drop n xs) = last xs"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1576
apply (induct xs arbitrary: n)
17501
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1577
 apply simp
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1578
apply (auto split:nat.split)
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1579
done
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1580
30128
365ee7319b86 revert some Suc 0 lemmas back to their original forms; added some simp rules for (1::nat)
huffman
parents: 30079
diff changeset
  1581
lemma last_conv_nth: "xs\<noteq>[] \<Longrightarrow> last xs = xs!(length xs - 1)"
17589
58eeffd73be1 renamed rules to iprover
nipkow
parents: 17501
diff changeset
  1582
by(induct xs)(auto simp:neq_Nil_conv)
58eeffd73be1 renamed rules to iprover
nipkow
parents: 17501
diff changeset
  1583
30128
365ee7319b86 revert some Suc 0 lemmas back to their original forms; added some simp rules for (1::nat)
huffman
parents: 30079
diff changeset
  1584
lemma butlast_conv_take: "butlast xs = take (length xs - 1) xs"
26584
46f3b89b2445 move lemmas from Word/BinBoolList.thy to List.thy
huffman
parents: 26480
diff changeset
  1585
by (induct xs, simp, case_tac xs, simp_all)
46f3b89b2445 move lemmas from Word/BinBoolList.thy to List.thy
huffman
parents: 26480
diff changeset
  1586
31077
28dd6fd3d184 more lemmas
nipkow
parents: 31055
diff changeset
  1587
lemma last_list_update:
28dd6fd3d184 more lemmas
nipkow
parents: 31055
diff changeset
  1588
  "xs \<noteq> [] \<Longrightarrow> last(xs[k:=x]) = (if k = size xs - 1 then x else last xs)"
28dd6fd3d184 more lemmas
nipkow
parents: 31055
diff changeset
  1589
by (auto simp: last_conv_nth)
28dd6fd3d184 more lemmas
nipkow
parents: 31055
diff changeset
  1590
28dd6fd3d184 more lemmas
nipkow
parents: 31055
diff changeset
  1591
lemma butlast_list_update:
28dd6fd3d184 more lemmas
nipkow
parents: 31055
diff changeset
  1592
  "butlast(xs[k:=x]) =
28dd6fd3d184 more lemmas
nipkow
parents: 31055
diff changeset
  1593
 (if k = size xs - 1 then butlast xs else (butlast xs)[k:=x])"
28dd6fd3d184 more lemmas
nipkow
parents: 31055
diff changeset
  1594
apply(cases xs rule:rev_cases)
28dd6fd3d184 more lemmas
nipkow
parents: 31055
diff changeset
  1595
apply simp
28dd6fd3d184 more lemmas
nipkow
parents: 31055
diff changeset
  1596
apply(simp add:list_update_append split:nat.splits)
28dd6fd3d184 more lemmas
nipkow
parents: 31055
diff changeset
  1597
done
28dd6fd3d184 more lemmas
nipkow
parents: 31055
diff changeset
  1598
36851
5135adb33157 added lemmas concerning last, butlast, insort
haftmann
parents: 36622
diff changeset
  1599
lemma last_map:
5135adb33157 added lemmas concerning last, butlast, insort
haftmann
parents: 36622
diff changeset
  1600
  "xs \<noteq> [] \<Longrightarrow> last (map f xs) = f (last xs)"
5135adb33157 added lemmas concerning last, butlast, insort
haftmann
parents: 36622
diff changeset
  1601
  by (cases xs rule: rev_cases) simp_all
5135adb33157 added lemmas concerning last, butlast, insort
haftmann
parents: 36622
diff changeset
  1602
5135adb33157 added lemmas concerning last, butlast, insort
haftmann
parents: 36622
diff changeset
  1603
lemma map_butlast:
5135adb33157 added lemmas concerning last, butlast, insort
haftmann
parents: 36622
diff changeset
  1604
  "map f (butlast xs) = butlast (map f xs)"
5135adb33157 added lemmas concerning last, butlast, insort
haftmann
parents: 36622
diff changeset
  1605
  by (induct xs) simp_all
5135adb33157 added lemmas concerning last, butlast, insort
haftmann
parents: 36622
diff changeset
  1606
40230
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  1607
lemma snoc_eq_iff_butlast:
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  1608
  "xs @ [x] = ys \<longleftrightarrow> (ys \<noteq> [] & butlast ys = xs & last ys = x)"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  1609
by (metis append_butlast_last_id append_is_Nil_conv butlast_snoc last_snoc not_Cons_self)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  1610
24796
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  1611
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15307
diff changeset
  1612
subsubsection {* @{text take} and @{text drop} *}
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1613
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  1614
lemma take_0 [simp]: "take 0 xs = []"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1615
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1616
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  1617
lemma drop_0 [simp]: "drop 0 xs = xs"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1618
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1619
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  1620
lemma take_Suc_Cons [simp]: "take (Suc n) (x # xs) = x # take n xs"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1621
by simp
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1622
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  1623
lemma drop_Suc_Cons [simp]: "drop (Suc n) (x # xs) = drop n xs"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1624
by simp
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1625
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  1626
declare take_Cons [simp del] and drop_Cons [simp del]
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1627
30128
365ee7319b86 revert some Suc 0 lemmas back to their original forms; added some simp rules for (1::nat)
huffman
parents: 30079
diff changeset
  1628
lemma take_1_Cons [simp]: "take 1 (x # xs) = [x]"
365ee7319b86 revert some Suc 0 lemmas back to their original forms; added some simp rules for (1::nat)
huffman
parents: 30079
diff changeset
  1629
  unfolding One_nat_def by simp
365ee7319b86 revert some Suc 0 lemmas back to their original forms; added some simp rules for (1::nat)
huffman
parents: 30079
diff changeset
  1630
365ee7319b86 revert some Suc 0 lemmas back to their original forms; added some simp rules for (1::nat)
huffman
parents: 30079
diff changeset
  1631
lemma drop_1_Cons [simp]: "drop 1 (x # xs) = xs"
365ee7319b86 revert some Suc 0 lemmas back to their original forms; added some simp rules for (1::nat)
huffman
parents: 30079
diff changeset
  1632
  unfolding One_nat_def by simp
365ee7319b86 revert some Suc 0 lemmas back to their original forms; added some simp rules for (1::nat)
huffman
parents: 30079
diff changeset
  1633
15110
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  1634
lemma take_Suc: "xs ~= [] ==> take (Suc n) xs = hd xs # take n (tl xs)"
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  1635
by(clarsimp simp add:neq_Nil_conv)
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  1636
14187
26dfcd0ac436 Added new theorems
nipkow
parents: 14111
diff changeset
  1637
lemma drop_Suc: "drop (Suc n) xs = drop n (tl xs)"
26dfcd0ac436 Added new theorems
nipkow
parents: 14111
diff changeset
  1638
by(cases xs, simp_all)
26dfcd0ac436 Added new theorems
nipkow
parents: 14111
diff changeset
  1639
26584
46f3b89b2445 move lemmas from Word/BinBoolList.thy to List.thy
huffman
parents: 26480
diff changeset
  1640
lemma take_tl: "take n (tl xs) = tl (take (Suc n) xs)"
46f3b89b2445 move lemmas from Word/BinBoolList.thy to List.thy
huffman
parents: 26480
diff changeset
  1641
by (induct xs arbitrary: n) simp_all
46f3b89b2445 move lemmas from Word/BinBoolList.thy to List.thy
huffman
parents: 26480
diff changeset
  1642
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1643
lemma drop_tl: "drop n (tl xs) = tl(drop n xs)"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1644
by(induct xs arbitrary: n, simp_all add:drop_Cons drop_Suc split:nat.split)
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1645
26584
46f3b89b2445 move lemmas from Word/BinBoolList.thy to List.thy
huffman
parents: 26480
diff changeset
  1646
lemma tl_take: "tl (take n xs) = take (n - 1) (tl xs)"
46f3b89b2445 move lemmas from Word/BinBoolList.thy to List.thy
huffman
parents: 26480
diff changeset
  1647
by (cases n, simp, cases xs, auto)
46f3b89b2445 move lemmas from Word/BinBoolList.thy to List.thy
huffman
parents: 26480
diff changeset
  1648
46f3b89b2445 move lemmas from Word/BinBoolList.thy to List.thy
huffman
parents: 26480
diff changeset
  1649
lemma tl_drop: "tl (drop n xs) = drop n (tl xs)"
46f3b89b2445 move lemmas from Word/BinBoolList.thy to List.thy
huffman
parents: 26480
diff changeset
  1650
by (simp only: drop_tl)
46f3b89b2445 move lemmas from Word/BinBoolList.thy to List.thy
huffman
parents: 26480
diff changeset
  1651
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1652
lemma nth_via_drop: "drop n xs = y#ys \<Longrightarrow> xs!n = y"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1653
apply (induct xs arbitrary: n, simp)
14187
26dfcd0ac436 Added new theorems
nipkow
parents: 14111
diff changeset
  1654
apply(simp add:drop_Cons nth_Cons split:nat.splits)
26dfcd0ac436 Added new theorems
nipkow
parents: 14111
diff changeset
  1655
done
26dfcd0ac436 Added new theorems
nipkow
parents: 14111
diff changeset
  1656
13913
b3ed67af04b8 Added take/dropWhile thms
nipkow
parents: 13883
diff changeset
  1657
lemma take_Suc_conv_app_nth:
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1658
  "i < length xs \<Longrightarrow> take (Suc i) xs = take i xs @ [xs!i]"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1659
apply (induct xs arbitrary: i, simp)
14208
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
  1660
apply (case_tac i, auto)
13913
b3ed67af04b8 Added take/dropWhile thms
nipkow
parents: 13883
diff changeset
  1661
done
b3ed67af04b8 Added take/dropWhile thms
nipkow
parents: 13883
diff changeset
  1662
14591
7be4d5dadf15 lemma drop_Suc_conv_tl added.
mehta
parents: 14589
diff changeset
  1663
lemma drop_Suc_conv_tl:
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1664
  "i < length xs \<Longrightarrow> (xs!i) # (drop (Suc i) xs) = drop i xs"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1665
apply (induct xs arbitrary: i, simp)
14591
7be4d5dadf15 lemma drop_Suc_conv_tl added.
mehta
parents: 14589
diff changeset
  1666
apply (case_tac i, auto)
7be4d5dadf15 lemma drop_Suc_conv_tl added.
mehta
parents: 14589
diff changeset
  1667
done
7be4d5dadf15 lemma drop_Suc_conv_tl added.
mehta
parents: 14589
diff changeset
  1668
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1669
lemma length_take [simp]: "length (take n xs) = min (length xs) n"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1670
by (induct n arbitrary: xs) (auto, case_tac xs, auto)
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1671
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1672
lemma length_drop [simp]: "length (drop n xs) = (length xs - n)"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1673
by (induct n arbitrary: xs) (auto, case_tac xs, auto)
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1674
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1675
lemma take_all [simp]: "length xs <= n ==> take n xs = xs"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1676
by (induct n arbitrary: xs) (auto, case_tac xs, auto)
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1677
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1678
lemma drop_all [simp]: "length xs <= n ==> drop n xs = []"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1679
by (induct n arbitrary: xs) (auto, case_tac xs, auto)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1680
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  1681
lemma take_append [simp]:
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1682
  "take n (xs @ ys) = (take n xs @ take (n - length xs) ys)"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1683
by (induct n arbitrary: xs) (auto, case_tac xs, auto)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1684
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  1685
lemma drop_append [simp]:
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1686
  "drop n (xs @ ys) = drop n xs @ drop (n - length xs) ys"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1687
by (induct n arbitrary: xs) (auto, case_tac xs, auto)
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1688
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1689
lemma take_take [simp]: "take n (take m xs) = take (min n m) xs"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1690
apply (induct m arbitrary: xs n, auto)
14208
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
  1691
apply (case_tac xs, auto)
15236
f289e8ba2bb3 Proofs needed to be updated because induction now preserves name of
nipkow
parents: 15176
diff changeset
  1692
apply (case_tac n, auto)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1693
done
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1694
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1695
lemma drop_drop [simp]: "drop n (drop m xs) = drop (n + m) xs"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1696
apply (induct m arbitrary: xs, auto)
14208
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
  1697
apply (case_tac xs, auto)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1698
done
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1699
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1700
lemma take_drop: "take n (drop m xs) = drop m (take (n + m) xs)"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1701
apply (induct m arbitrary: xs n, auto)
14208
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
  1702
apply (case_tac xs, auto)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1703
done
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1704
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1705
lemma drop_take: "drop n (take m xs) = take (m-n) (drop n xs)"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1706
apply(induct xs arbitrary: m n)
14802
e05116289ff9 added drop_take:thm
nipkow
parents: 14770
diff changeset
  1707
 apply simp
e05116289ff9 added drop_take:thm
nipkow
parents: 14770
diff changeset
  1708
apply(simp add: take_Cons drop_Cons split:nat.split)
e05116289ff9 added drop_take:thm
nipkow
parents: 14770
diff changeset
  1709
done
e05116289ff9 added drop_take:thm
nipkow
parents: 14770
diff changeset
  1710
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1711
lemma append_take_drop_id [simp]: "take n xs @ drop n xs = xs"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1712
apply (induct n arbitrary: xs, auto)
14208
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
  1713
apply (case_tac xs, auto)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1714
done
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1715
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1716
lemma take_eq_Nil[simp]: "(take n xs = []) = (n = 0 \<or> xs = [])"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1717
apply(induct xs arbitrary: n)
15110
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  1718
 apply simp
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  1719
apply(simp add:take_Cons split:nat.split)
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  1720
done
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  1721
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1722
lemma drop_eq_Nil[simp]: "(drop n xs = []) = (length xs <= n)"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1723
apply(induct xs arbitrary: n)
15110
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  1724
apply simp
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  1725
apply(simp add:drop_Cons split:nat.split)
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  1726
done
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  1727
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1728
lemma take_map: "take n (map f xs) = map f (take n xs)"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1729
apply (induct n arbitrary: xs, auto)
14208
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
  1730
apply (case_tac xs, auto)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1731
done
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1732
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1733
lemma drop_map: "drop n (map f xs) = map f (drop n xs)"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1734
apply (induct n arbitrary: xs, auto)
14208
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
  1735
apply (case_tac xs, auto)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1736
done
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1737
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1738
lemma rev_take: "rev (take i xs) = drop (length xs - i) (rev xs)"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1739
apply (induct xs arbitrary: i, auto)
14208
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
  1740
apply (case_tac i, auto)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1741
done
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1742
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1743
lemma rev_drop: "rev (drop i xs) = take (length xs - i) (rev xs)"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1744
apply (induct xs arbitrary: i, auto)
14208
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
  1745
apply (case_tac i, auto)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1746
done
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1747
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1748
lemma nth_take [simp]: "i < n ==> (take n xs)!i = xs!i"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1749
apply (induct xs arbitrary: i n, auto)
14208
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
  1750
apply (case_tac n, blast)
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
  1751
apply (case_tac i, auto)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1752
done
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1753
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  1754
lemma nth_drop [simp]:
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1755
  "n + i <= length xs ==> (drop n xs)!i = xs!(n + i)"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1756
apply (induct n arbitrary: xs i, auto)
14208
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
  1757
apply (case_tac xs, auto)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1758
done
3507
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
  1759
26584
46f3b89b2445 move lemmas from Word/BinBoolList.thy to List.thy
huffman
parents: 26480
diff changeset
  1760
lemma butlast_take:
30128
365ee7319b86 revert some Suc 0 lemmas back to their original forms; added some simp rules for (1::nat)
huffman
parents: 30079
diff changeset
  1761
  "n <= length xs ==> butlast (take n xs) = take (n - 1) xs"
26584
46f3b89b2445 move lemmas from Word/BinBoolList.thy to List.thy
huffman
parents: 26480
diff changeset
  1762
by (simp add: butlast_conv_take min_max.inf_absorb1 min_max.inf_absorb2)
46f3b89b2445 move lemmas from Word/BinBoolList.thy to List.thy
huffman
parents: 26480
diff changeset
  1763
46f3b89b2445 move lemmas from Word/BinBoolList.thy to List.thy
huffman
parents: 26480
diff changeset
  1764
lemma butlast_drop: "butlast (drop n xs) = drop n (butlast xs)"
30128
365ee7319b86 revert some Suc 0 lemmas back to their original forms; added some simp rules for (1::nat)
huffman
parents: 30079
diff changeset
  1765
by (simp add: butlast_conv_take drop_take add_ac)
26584
46f3b89b2445 move lemmas from Word/BinBoolList.thy to List.thy
huffman
parents: 26480
diff changeset
  1766
46f3b89b2445 move lemmas from Word/BinBoolList.thy to List.thy
huffman
parents: 26480
diff changeset
  1767
lemma take_butlast: "n < length xs ==> take n (butlast xs) = take n xs"
46f3b89b2445 move lemmas from Word/BinBoolList.thy to List.thy
huffman
parents: 26480
diff changeset
  1768
by (simp add: butlast_conv_take min_max.inf_absorb1)
46f3b89b2445 move lemmas from Word/BinBoolList.thy to List.thy
huffman
parents: 26480
diff changeset
  1769
46f3b89b2445 move lemmas from Word/BinBoolList.thy to List.thy
huffman
parents: 26480
diff changeset
  1770
lemma drop_butlast: "drop n (butlast xs) = butlast (drop n xs)"
30128
365ee7319b86 revert some Suc 0 lemmas back to their original forms; added some simp rules for (1::nat)
huffman
parents: 30079
diff changeset
  1771
by (simp add: butlast_conv_take drop_take add_ac)
26584
46f3b89b2445 move lemmas from Word/BinBoolList.thy to List.thy
huffman
parents: 26480
diff changeset
  1772
18423
d7859164447f new lemmas
nipkow
parents: 18336
diff changeset
  1773
lemma hd_drop_conv_nth: "\<lbrakk> xs \<noteq> []; n < length xs \<rbrakk> \<Longrightarrow> hd(drop n xs) = xs!n"
d7859164447f new lemmas
nipkow
parents: 18336
diff changeset
  1774
by(simp add: hd_conv_nth)
d7859164447f new lemmas
nipkow
parents: 18336
diff changeset
  1775
35248
e64950874224 added lemma
nipkow
parents: 35217
diff changeset
  1776
lemma set_take_subset_set_take:
e64950874224 added lemma
nipkow
parents: 35217
diff changeset
  1777
  "m <= n \<Longrightarrow> set(take m xs) <= set(take n xs)"
41463
edbf0a86fb1c adding simproc to rewrite list comprehensions to set comprehensions; adopting proofs
bulwahn
parents: 41372
diff changeset
  1778
apply (induct xs arbitrary: m n)
edbf0a86fb1c adding simproc to rewrite list comprehensions to set comprehensions; adopting proofs
bulwahn
parents: 41372
diff changeset
  1779
apply simp
edbf0a86fb1c adding simproc to rewrite list comprehensions to set comprehensions; adopting proofs
bulwahn
parents: 41372
diff changeset
  1780
apply (case_tac n)
edbf0a86fb1c adding simproc to rewrite list comprehensions to set comprehensions; adopting proofs
bulwahn
parents: 41372
diff changeset
  1781
apply (auto simp: take_Cons)
edbf0a86fb1c adding simproc to rewrite list comprehensions to set comprehensions; adopting proofs
bulwahn
parents: 41372
diff changeset
  1782
done
35248
e64950874224 added lemma
nipkow
parents: 35217
diff changeset
  1783
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1784
lemma set_take_subset: "set(take n xs) \<subseteq> set xs"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1785
by(induct xs arbitrary: n)(auto simp:take_Cons split:nat.split)
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1786
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1787
lemma set_drop_subset: "set(drop n xs) \<subseteq> set xs"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1788
by(induct xs arbitrary: n)(auto simp:drop_Cons split:nat.split)
14025
d9b155757dc8 *** empty log message ***
nipkow
parents: 13913
diff changeset
  1789
35248
e64950874224 added lemma
nipkow
parents: 35217
diff changeset
  1790
lemma set_drop_subset_set_drop:
e64950874224 added lemma
nipkow
parents: 35217
diff changeset
  1791
  "m >= n \<Longrightarrow> set(drop m xs) <= set(drop n xs)"
e64950874224 added lemma
nipkow
parents: 35217
diff changeset
  1792
apply(induct xs arbitrary: m n)
e64950874224 added lemma
nipkow
parents: 35217
diff changeset
  1793
apply(auto simp:drop_Cons split:nat.split)
e64950874224 added lemma
nipkow
parents: 35217
diff changeset
  1794
apply (metis set_drop_subset subset_iff)
e64950874224 added lemma
nipkow
parents: 35217
diff changeset
  1795
done
e64950874224 added lemma
nipkow
parents: 35217
diff changeset
  1796
14187
26dfcd0ac436 Added new theorems
nipkow
parents: 14111
diff changeset
  1797
lemma in_set_takeD: "x : set(take n xs) \<Longrightarrow> x : set xs"
26dfcd0ac436 Added new theorems
nipkow
parents: 14111
diff changeset
  1798
using set_take_subset by fast
26dfcd0ac436 Added new theorems
nipkow
parents: 14111
diff changeset
  1799
26dfcd0ac436 Added new theorems
nipkow
parents: 14111
diff changeset
  1800
lemma in_set_dropD: "x : set(drop n xs) \<Longrightarrow> x : set xs"
26dfcd0ac436 Added new theorems
nipkow
parents: 14111
diff changeset
  1801
using set_drop_subset by fast
26dfcd0ac436 Added new theorems
nipkow
parents: 14111
diff changeset
  1802
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1803
lemma append_eq_conv_conj:
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1804
  "(xs @ ys = zs) = (xs = take (length xs) zs \<and> ys = drop (length xs) zs)"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1805
apply (induct xs arbitrary: zs, simp, clarsimp)
14208
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
  1806
apply (case_tac zs, auto)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1807
done
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  1808
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1809
lemma take_add: 
42713
276c8cbeb5d2 removed assumption from lemma List.take_add
noschinl
parents: 42411
diff changeset
  1810
  "take (i+j) xs = take i xs @ take j (drop i xs)"
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1811
apply (induct xs arbitrary: i, auto) 
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1812
apply (case_tac i, simp_all)
14050
826037db30cd new theorem
paulson
parents: 14025
diff changeset
  1813
done
826037db30cd new theorem
paulson
parents: 14025
diff changeset
  1814
14300
bf8b8c9425c3 *** empty log message ***
nipkow
parents: 14247
diff changeset
  1815
lemma append_eq_append_conv_if:
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1816
 "(xs\<^isub>1 @ xs\<^isub>2 = ys\<^isub>1 @ ys\<^isub>2) =
14300
bf8b8c9425c3 *** empty log message ***
nipkow
parents: 14247
diff changeset
  1817
  (if size xs\<^isub>1 \<le> size ys\<^isub>1
bf8b8c9425c3 *** empty log message ***
nipkow
parents: 14247
diff changeset
  1818
   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
bf8b8c9425c3 *** empty log message ***
nipkow
parents: 14247
diff changeset
  1819
   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)"
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1820
apply(induct xs\<^isub>1 arbitrary: ys\<^isub>1)
14300
bf8b8c9425c3 *** empty log message ***
nipkow
parents: 14247
diff changeset
  1821
 apply simp
bf8b8c9425c3 *** empty log message ***
nipkow
parents: 14247
diff changeset
  1822
apply(case_tac ys\<^isub>1)
bf8b8c9425c3 *** empty log message ***
nipkow
parents: 14247
diff changeset
  1823
apply simp_all
bf8b8c9425c3 *** empty log message ***
nipkow
parents: 14247
diff changeset
  1824
done
bf8b8c9425c3 *** empty log message ***
nipkow
parents: 14247
diff changeset
  1825
15110
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  1826
lemma take_hd_drop:
30079
293b896b9c25 make proofs work whether or not One_nat_def is a simp rule; replace 1 with Suc 0 in the rhs of some simp rules
huffman
parents: 30008
diff changeset
  1827
  "n < length xs \<Longrightarrow> take n xs @ [hd (drop n xs)] = take (Suc n) xs"
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  1828
apply(induct xs arbitrary: n)
15110
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  1829
apply simp
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  1830
apply(simp add:drop_Cons split:nat.split)
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  1831
done
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  1832
17501
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1833
lemma id_take_nth_drop:
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1834
 "i < length xs \<Longrightarrow> xs = take i xs @ xs!i # drop (Suc i) xs" 
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1835
proof -
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1836
  assume si: "i < length xs"
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1837
  hence "xs = take (Suc i) xs @ drop (Suc i) xs" by auto
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1838
  moreover
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1839
  from si have "take (Suc i) xs = take i xs @ [xs!i]"
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1840
    apply (rule_tac take_Suc_conv_app_nth) by arith
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1841
  ultimately show ?thesis by auto
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1842
qed
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1843
  
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1844
lemma upd_conv_take_nth_drop:
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1845
 "i < length xs \<Longrightarrow> xs[i:=a] = take i xs @ a # drop (Suc i) xs"
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1846
proof -
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1847
  assume i: "i < length xs"
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1848
  have "xs[i:=a] = (take i xs @ xs!i # drop (Suc i) xs)[i:=a]"
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1849
    by(rule arg_cong[OF id_take_nth_drop[OF i]])
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1850
  also have "\<dots> = take i xs @ a # drop (Suc i) xs"
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1851
    using i by (simp add: list_update_append)
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1852
  finally show ?thesis .
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1853
qed
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1854
24796
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  1855
lemma nth_drop':
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  1856
  "i < length xs \<Longrightarrow> xs ! i # drop (Suc i) xs = drop i xs"
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  1857
apply (induct i arbitrary: xs)
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  1858
apply (simp add: neq_Nil_conv)
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  1859
apply (erule exE)+
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  1860
apply simp
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  1861
apply (case_tac xs)
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  1862
apply simp_all
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  1863
done
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  1864
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1865
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15307
diff changeset
  1866
subsubsection {* @{text takeWhile} and @{text dropWhile} *}
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1867
33639
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1868
lemma length_takeWhile_le: "length (takeWhile P xs) \<le> length xs"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1869
  by (induct xs) auto
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1870
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  1871
lemma takeWhile_dropWhile_id [simp]: "takeWhile P xs @ dropWhile P xs = xs"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1872
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1873
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  1874
lemma takeWhile_append1 [simp]:
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1875
"[| x:set xs; ~P(x)|] ==> takeWhile P (xs @ ys) = takeWhile P xs"
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1876
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1877
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  1878
lemma takeWhile_append2 [simp]:
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1879
"(!!x. x : set xs ==> P x) ==> takeWhile P (xs @ ys) = xs @ takeWhile P ys"
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1880
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1881
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  1882
lemma takeWhile_tail: "\<not> P x ==> takeWhile P (xs @ (x#l)) = takeWhile P xs"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1883
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1884
33639
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1885
lemma takeWhile_nth: "j < length (takeWhile P xs) \<Longrightarrow> takeWhile P xs ! j = xs ! j"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1886
apply (subst (3) takeWhile_dropWhile_id[symmetric]) unfolding nth_append by auto
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1887
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1888
lemma dropWhile_nth: "j < length (dropWhile P xs) \<Longrightarrow> dropWhile P xs ! j = xs ! (j + length (takeWhile P xs))"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1889
apply (subst (3) takeWhile_dropWhile_id[symmetric]) unfolding nth_append by auto
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1890
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1891
lemma length_dropWhile_le: "length (dropWhile P xs) \<le> length xs"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1892
by (induct xs) auto
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1893
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  1894
lemma dropWhile_append1 [simp]:
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1895
"[| x : set xs; ~P(x)|] ==> dropWhile P (xs @ ys) = (dropWhile P xs)@ys"
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1896
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1897
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  1898
lemma dropWhile_append2 [simp]:
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1899
"(!!x. x:set xs ==> P(x)) ==> dropWhile P (xs @ ys) = dropWhile P ys"
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1900
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1901
23971
e6d505d5b03d renamed lemma "set_take_whileD" to "set_takeWhileD"
krauss
parents: 23740
diff changeset
  1902
lemma set_takeWhileD: "x : set (takeWhile P xs) ==> x : set xs \<and> P x"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  1903
by (induct xs) (auto split: split_if_asm)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  1904
13913
b3ed67af04b8 Added take/dropWhile thms
nipkow
parents: 13883
diff changeset
  1905
lemma takeWhile_eq_all_conv[simp]:
b3ed67af04b8 Added take/dropWhile thms
nipkow
parents: 13883
diff changeset
  1906
 "(takeWhile P xs = xs) = (\<forall>x \<in> set xs. P x)"
b3ed67af04b8 Added take/dropWhile thms
nipkow
parents: 13883
diff changeset
  1907
by(induct xs, auto)
b3ed67af04b8 Added take/dropWhile thms
nipkow
parents: 13883
diff changeset
  1908
b3ed67af04b8 Added take/dropWhile thms
nipkow
parents: 13883
diff changeset
  1909
lemma dropWhile_eq_Nil_conv[simp]:
b3ed67af04b8 Added take/dropWhile thms
nipkow
parents: 13883
diff changeset
  1910
 "(dropWhile P xs = []) = (\<forall>x \<in> set xs. P x)"
b3ed67af04b8 Added take/dropWhile thms
nipkow
parents: 13883
diff changeset
  1911
by(induct xs, auto)
b3ed67af04b8 Added take/dropWhile thms
nipkow
parents: 13883
diff changeset
  1912
b3ed67af04b8 Added take/dropWhile thms
nipkow
parents: 13883
diff changeset
  1913
lemma dropWhile_eq_Cons_conv:
b3ed67af04b8 Added take/dropWhile thms
nipkow
parents: 13883
diff changeset
  1914
 "(dropWhile P xs = y#ys) = (xs = takeWhile P xs @ y # ys & \<not> P y)"
b3ed67af04b8 Added take/dropWhile thms
nipkow
parents: 13883
diff changeset
  1915
by(induct xs, auto)
b3ed67af04b8 Added take/dropWhile thms
nipkow
parents: 13883
diff changeset
  1916
31077
28dd6fd3d184 more lemmas
nipkow
parents: 31055
diff changeset
  1917
lemma distinct_takeWhile[simp]: "distinct xs ==> distinct (takeWhile P xs)"
28dd6fd3d184 more lemmas
nipkow
parents: 31055
diff changeset
  1918
by (induct xs) (auto dest: set_takeWhileD)
28dd6fd3d184 more lemmas
nipkow
parents: 31055
diff changeset
  1919
28dd6fd3d184 more lemmas
nipkow
parents: 31055
diff changeset
  1920
lemma distinct_dropWhile[simp]: "distinct xs ==> distinct (dropWhile P xs)"
28dd6fd3d184 more lemmas
nipkow
parents: 31055
diff changeset
  1921
by (induct xs) auto
28dd6fd3d184 more lemmas
nipkow
parents: 31055
diff changeset
  1922
33639
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1923
lemma takeWhile_map: "takeWhile P (map f xs) = map f (takeWhile (P \<circ> f) xs)"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1924
by (induct xs) auto
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1925
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1926
lemma dropWhile_map: "dropWhile P (map f xs) = map f (dropWhile (P \<circ> f) xs)"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1927
by (induct xs) auto
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1928
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1929
lemma takeWhile_eq_take: "takeWhile P xs = take (length (takeWhile P xs)) xs"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1930
by (induct xs) auto
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1931
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1932
lemma dropWhile_eq_drop: "dropWhile P xs = drop (length (takeWhile P xs)) xs"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1933
by (induct xs) auto
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1934
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1935
lemma hd_dropWhile:
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1936
  "dropWhile P xs \<noteq> [] \<Longrightarrow> \<not> P (hd (dropWhile P xs))"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1937
using assms by (induct xs) auto
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1938
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1939
lemma takeWhile_eq_filter:
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1940
  assumes "\<And> x. x \<in> set (dropWhile P xs) \<Longrightarrow> \<not> P x"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1941
  shows "takeWhile P xs = filter P xs"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1942
proof -
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1943
  have A: "filter P xs = filter P (takeWhile P xs @ dropWhile P xs)"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1944
    by simp
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1945
  have B: "filter P (dropWhile P xs) = []"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1946
    unfolding filter_empty_conv using assms by blast
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1947
  have "filter P xs = takeWhile P xs"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1948
    unfolding A filter_append B
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1949
    by (auto simp add: filter_id_conv dest: set_takeWhileD)
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1950
  thus ?thesis ..
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1951
qed
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1952
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1953
lemma takeWhile_eq_take_P_nth:
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1954
  "\<lbrakk> \<And> i. \<lbrakk> i < n ; i < length xs \<rbrakk> \<Longrightarrow> P (xs ! i) ; n < length xs \<Longrightarrow> \<not> P (xs ! n) \<rbrakk> \<Longrightarrow>
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1955
  takeWhile P xs = take n xs"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1956
proof (induct xs arbitrary: n)
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1957
  case (Cons x xs)
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1958
  thus ?case
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1959
  proof (cases n)
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1960
    case (Suc n') note this[simp]
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1961
    have "P x" using Cons.prems(1)[of 0] by simp
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1962
    moreover have "takeWhile P xs = take n' xs"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1963
    proof (rule Cons.hyps)
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1964
      case goal1 thus "P (xs ! i)" using Cons.prems(1)[of "Suc i"] by simp
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1965
    next case goal2 thus ?case using Cons by auto
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1966
    qed
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1967
    ultimately show ?thesis by simp
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1968
   qed simp
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1969
qed simp
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1970
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1971
lemma nth_length_takeWhile:
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1972
  "length (takeWhile P xs) < length xs \<Longrightarrow> \<not> P (xs ! length (takeWhile P xs))"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1973
by (induct xs) auto
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1974
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1975
lemma length_takeWhile_less_P_nth:
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1976
  assumes all: "\<And> i. i < j \<Longrightarrow> P (xs ! i)" and "j \<le> length xs"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1977
  shows "j \<le> length (takeWhile P xs)"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1978
proof (rule classical)
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1979
  assume "\<not> ?thesis"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1980
  hence "length (takeWhile P xs) < length xs" using assms by simp
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1981
  thus ?thesis using all `\<not> ?thesis` nth_length_takeWhile[of P xs] by auto
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  1982
qed
31077
28dd6fd3d184 more lemmas
nipkow
parents: 31055
diff changeset
  1983
17501
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1984
text{* The following two lemmmas could be generalized to an arbitrary
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1985
property. *}
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1986
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1987
lemma takeWhile_neq_rev: "\<lbrakk>distinct xs; x \<in> set xs\<rbrakk> \<Longrightarrow>
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1988
 takeWhile (\<lambda>y. y \<noteq> x) (rev xs) = rev (tl (dropWhile (\<lambda>y. y \<noteq> x) xs))"
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1989
by(induct xs) (auto simp: takeWhile_tail[where l="[]"])
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1990
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1991
lemma dropWhile_neq_rev: "\<lbrakk>distinct xs; x \<in> set xs\<rbrakk> \<Longrightarrow>
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1992
  dropWhile (\<lambda>y. y \<noteq> x) (rev xs) = x # rev (takeWhile (\<lambda>y. y \<noteq> x) xs)"
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1993
apply(induct xs)
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1994
 apply simp
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1995
apply auto
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1996
apply(subst dropWhile_append2)
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1997
apply auto
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1998
done
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  1999
18423
d7859164447f new lemmas
nipkow
parents: 18336
diff changeset
  2000
lemma takeWhile_not_last:
d7859164447f new lemmas
nipkow
parents: 18336
diff changeset
  2001
 "\<lbrakk> xs \<noteq> []; distinct xs\<rbrakk> \<Longrightarrow> takeWhile (\<lambda>y. y \<noteq> last xs) xs = butlast xs"
d7859164447f new lemmas
nipkow
parents: 18336
diff changeset
  2002
apply(induct xs)
d7859164447f new lemmas
nipkow
parents: 18336
diff changeset
  2003
 apply simp
d7859164447f new lemmas
nipkow
parents: 18336
diff changeset
  2004
apply(case_tac xs)
d7859164447f new lemmas
nipkow
parents: 18336
diff changeset
  2005
apply(auto)
d7859164447f new lemmas
nipkow
parents: 18336
diff changeset
  2006
done
d7859164447f new lemmas
nipkow
parents: 18336
diff changeset
  2007
44013
5cfc1c36ae97 moved recdef package to HOL/Library/Old_Recdef.thy
krauss
parents: 43594
diff changeset
  2008
lemma takeWhile_cong [fundef_cong]:
18336
1a2e30b37ed3 Added recdef congruence rules for bounded quantifiers and commonly used
krauss
parents: 18049
diff changeset
  2009
  "[| l = k; !!x. x : set l ==> P x = Q x |] 
1a2e30b37ed3 Added recdef congruence rules for bounded quantifiers and commonly used
krauss
parents: 18049
diff changeset
  2010
  ==> takeWhile P l = takeWhile Q k"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  2011
by (induct k arbitrary: l) (simp_all)
18336
1a2e30b37ed3 Added recdef congruence rules for bounded quantifiers and commonly used
krauss
parents: 18049
diff changeset
  2012
44013
5cfc1c36ae97 moved recdef package to HOL/Library/Old_Recdef.thy
krauss
parents: 43594
diff changeset
  2013
lemma dropWhile_cong [fundef_cong]:
18336
1a2e30b37ed3 Added recdef congruence rules for bounded quantifiers and commonly used
krauss
parents: 18049
diff changeset
  2014
  "[| l = k; !!x. x : set l ==> P x = Q x |] 
1a2e30b37ed3 Added recdef congruence rules for bounded quantifiers and commonly used
krauss
parents: 18049
diff changeset
  2015
  ==> dropWhile P l = dropWhile Q k"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  2016
by (induct k arbitrary: l, simp_all)
18336
1a2e30b37ed3 Added recdef congruence rules for bounded quantifiers and commonly used
krauss
parents: 18049
diff changeset
  2017
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2018
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15307
diff changeset
  2019
subsubsection {* @{text zip} *}
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2020
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  2021
lemma zip_Nil [simp]: "zip [] ys = []"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2022
by (induct ys) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2023
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  2024
lemma zip_Cons_Cons [simp]: "zip (x # xs) (y # ys) = (x, y) # zip xs ys"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2025
by simp
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2026
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  2027
declare zip_Cons [simp del]
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2028
36198
ead2db2be11a more convenient equations for zip
haftmann
parents: 35828
diff changeset
  2029
lemma [code]:
ead2db2be11a more convenient equations for zip
haftmann
parents: 35828
diff changeset
  2030
  "zip [] ys = []"
ead2db2be11a more convenient equations for zip
haftmann
parents: 35828
diff changeset
  2031
  "zip xs [] = []"
ead2db2be11a more convenient equations for zip
haftmann
parents: 35828
diff changeset
  2032
  "zip (x # xs) (y # ys) = (x, y) # zip xs ys"
ead2db2be11a more convenient equations for zip
haftmann
parents: 35828
diff changeset
  2033
  by (fact zip_Nil zip.simps(1) zip_Cons_Cons)+
ead2db2be11a more convenient equations for zip
haftmann
parents: 35828
diff changeset
  2034
15281
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  2035
lemma zip_Cons1:
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  2036
 "zip (x#xs) ys = (case ys of [] \<Rightarrow> [] | y#ys \<Rightarrow> (x,y)#zip xs ys)"
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  2037
by(auto split:list.split)
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  2038
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  2039
lemma length_zip [simp]:
22493
db930e490fe5 added another rule for simultaneous induction, and lemmas for zip
krauss
parents: 22422
diff changeset
  2040
"length (zip xs ys) = min (length xs) (length ys)"
db930e490fe5 added another rule for simultaneous induction, and lemmas for zip
krauss
parents: 22422
diff changeset
  2041
by (induct xs ys rule:list_induct2') auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2042
34978
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  2043
lemma zip_obtain_same_length:
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  2044
  assumes "\<And>zs ws n. length zs = length ws \<Longrightarrow> n = min (length xs) (length ys)
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  2045
    \<Longrightarrow> zs = take n xs \<Longrightarrow> ws = take n ys \<Longrightarrow> P (zip zs ws)"
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  2046
  shows "P (zip xs ys)"
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  2047
proof -
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  2048
  let ?n = "min (length xs) (length ys)"
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  2049
  have "P (zip (take ?n xs) (take ?n ys))"
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  2050
    by (rule assms) simp_all
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  2051
  moreover have "zip xs ys = zip (take ?n xs) (take ?n ys)"
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  2052
  proof (induct xs arbitrary: ys)
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  2053
    case Nil then show ?case by simp
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  2054
  next
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  2055
    case (Cons x xs) then show ?case by (cases ys) simp_all
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  2056
  qed
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  2057
  ultimately show ?thesis by simp
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  2058
qed
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  2059
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2060
lemma zip_append1:
22493
db930e490fe5 added another rule for simultaneous induction, and lemmas for zip
krauss
parents: 22422
diff changeset
  2061
"zip (xs @ ys) zs =
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2062
zip xs (take (length xs) zs) @ zip ys (drop (length xs) zs)"
22493
db930e490fe5 added another rule for simultaneous induction, and lemmas for zip
krauss
parents: 22422
diff changeset
  2063
by (induct xs zs rule:list_induct2') auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2064
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2065
lemma zip_append2:
22493
db930e490fe5 added another rule for simultaneous induction, and lemmas for zip
krauss
parents: 22422
diff changeset
  2066
"zip xs (ys @ zs) =
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2067
zip (take (length ys) xs) ys @ zip (drop (length ys) xs) zs"
22493
db930e490fe5 added another rule for simultaneous induction, and lemmas for zip
krauss
parents: 22422
diff changeset
  2068
by (induct xs ys rule:list_induct2') auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2069
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  2070
lemma zip_append [simp]:
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  2071
 "[| length xs = length us; length ys = length vs |] ==>
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2072
zip (xs@ys) (us@vs) = zip xs us @ zip ys vs"
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2073
by (simp add: zip_append1)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2074
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2075
lemma zip_rev:
14247
cb32eb89bddd *** empty log message ***
nipkow
parents: 14208
diff changeset
  2076
"length xs = length ys ==> zip (rev xs) (rev ys) = rev (zip xs ys)"
cb32eb89bddd *** empty log message ***
nipkow
parents: 14208
diff changeset
  2077
by (induct rule:list_induct2, simp_all)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2078
33639
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2079
lemma zip_map_map:
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2080
  "zip (map f xs) (map g ys) = map (\<lambda> (x, y). (f x, g y)) (zip xs ys)"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2081
proof (induct xs arbitrary: ys)
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2082
  case (Cons x xs) note Cons_x_xs = Cons.hyps
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2083
  show ?case
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2084
  proof (cases ys)
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2085
    case (Cons y ys')
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2086
    show ?thesis unfolding Cons using Cons_x_xs by simp
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2087
  qed simp
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2088
qed simp
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2089
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2090
lemma zip_map1:
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2091
  "zip (map f xs) ys = map (\<lambda>(x, y). (f x, y)) (zip xs ys)"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2092
using zip_map_map[of f xs "\<lambda>x. x" ys] by simp
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2093
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2094
lemma zip_map2:
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2095
  "zip xs (map f ys) = map (\<lambda>(x, y). (x, f y)) (zip xs ys)"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2096
using zip_map_map[of "\<lambda>x. x" xs f ys] by simp
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2097
23096
423ad2fe9f76 *** empty log message ***
nipkow
parents: 23060
diff changeset
  2098
lemma map_zip_map:
33639
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2099
  "map f (zip (map g xs) ys) = map (%(x,y). f(g x, y)) (zip xs ys)"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2100
unfolding zip_map1 by auto
23096
423ad2fe9f76 *** empty log message ***
nipkow
parents: 23060
diff changeset
  2101
423ad2fe9f76 *** empty log message ***
nipkow
parents: 23060
diff changeset
  2102
lemma map_zip_map2:
33639
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2103
  "map f (zip xs (map g ys)) = map (%(x,y). f(x, g y)) (zip xs ys)"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2104
unfolding zip_map2 by auto
23096
423ad2fe9f76 *** empty log message ***
nipkow
parents: 23060
diff changeset
  2105
31080
21ffc770ebc0 lemmas by Andreas Lochbihler
nipkow
parents: 31077
diff changeset
  2106
text{* Courtesy of Andreas Lochbihler: *}
21ffc770ebc0 lemmas by Andreas Lochbihler
nipkow
parents: 31077
diff changeset
  2107
lemma zip_same_conv_map: "zip xs xs = map (\<lambda>x. (x, x)) xs"
21ffc770ebc0 lemmas by Andreas Lochbihler
nipkow
parents: 31077
diff changeset
  2108
by(induct xs) auto
21ffc770ebc0 lemmas by Andreas Lochbihler
nipkow
parents: 31077
diff changeset
  2109
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  2110
lemma nth_zip [simp]:
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2111
"[| i < length xs; i < length ys|] ==> (zip xs ys)!i = (xs!i, ys!i)"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2112
apply (induct ys arbitrary: i xs, simp)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2113
apply (case_tac xs)
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2114
 apply (simp_all add: nth.simps split: nat.split)
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2115
done
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2116
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2117
lemma set_zip:
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2118
"set (zip xs ys) = {(xs!i, ys!i) | i. i < min (length xs) (length ys)}"
31080
21ffc770ebc0 lemmas by Andreas Lochbihler
nipkow
parents: 31077
diff changeset
  2119
by(simp add: set_conv_nth cong: rev_conj_cong)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2120
33639
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2121
lemma zip_same: "((a,b) \<in> set (zip xs xs)) = (a \<in> set xs \<and> a = b)"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2122
by(induct xs) auto
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2123
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2124
lemma zip_update:
31080
21ffc770ebc0 lemmas by Andreas Lochbihler
nipkow
parents: 31077
diff changeset
  2125
  "zip (xs[i:=x]) (ys[i:=y]) = (zip xs ys)[i:=(x,y)]"
21ffc770ebc0 lemmas by Andreas Lochbihler
nipkow
parents: 31077
diff changeset
  2126
by(rule sym, simp add: update_zip)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2127
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  2128
lemma zip_replicate [simp]:
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2129
  "zip (replicate i x) (replicate j y) = replicate (min i j) (x,y)"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2130
apply (induct i arbitrary: j, auto)
14208
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
  2131
apply (case_tac j, auto)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2132
done
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2133
19487
d5e79a41bce0 added zip/take/drop lemmas
nipkow
parents: 19390
diff changeset
  2134
lemma take_zip:
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2135
  "take n (zip xs ys) = zip (take n xs) (take n ys)"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2136
apply (induct n arbitrary: xs ys)
19487
d5e79a41bce0 added zip/take/drop lemmas
nipkow
parents: 19390
diff changeset
  2137
 apply simp
d5e79a41bce0 added zip/take/drop lemmas
nipkow
parents: 19390
diff changeset
  2138
apply (case_tac xs, simp)
d5e79a41bce0 added zip/take/drop lemmas
nipkow
parents: 19390
diff changeset
  2139
apply (case_tac ys, simp_all)
d5e79a41bce0 added zip/take/drop lemmas
nipkow
parents: 19390
diff changeset
  2140
done
d5e79a41bce0 added zip/take/drop lemmas
nipkow
parents: 19390
diff changeset
  2141
d5e79a41bce0 added zip/take/drop lemmas
nipkow
parents: 19390
diff changeset
  2142
lemma drop_zip:
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2143
  "drop n (zip xs ys) = zip (drop n xs) (drop n ys)"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2144
apply (induct n arbitrary: xs ys)
19487
d5e79a41bce0 added zip/take/drop lemmas
nipkow
parents: 19390
diff changeset
  2145
 apply simp
d5e79a41bce0 added zip/take/drop lemmas
nipkow
parents: 19390
diff changeset
  2146
apply (case_tac xs, simp)
d5e79a41bce0 added zip/take/drop lemmas
nipkow
parents: 19390
diff changeset
  2147
apply (case_tac ys, simp_all)
d5e79a41bce0 added zip/take/drop lemmas
nipkow
parents: 19390
diff changeset
  2148
done
d5e79a41bce0 added zip/take/drop lemmas
nipkow
parents: 19390
diff changeset
  2149
33639
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2150
lemma zip_takeWhile_fst: "zip (takeWhile P xs) ys = takeWhile (P \<circ> fst) (zip xs ys)"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2151
proof (induct xs arbitrary: ys)
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2152
  case (Cons x xs) thus ?case by (cases ys) auto
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2153
qed simp
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2154
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2155
lemma zip_takeWhile_snd: "zip xs (takeWhile P ys) = takeWhile (P \<circ> snd) (zip xs ys)"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2156
proof (induct xs arbitrary: ys)
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2157
  case (Cons x xs) thus ?case by (cases ys) auto
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2158
qed simp
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2159
22493
db930e490fe5 added another rule for simultaneous induction, and lemmas for zip
krauss
parents: 22422
diff changeset
  2160
lemma set_zip_leftD:
db930e490fe5 added another rule for simultaneous induction, and lemmas for zip
krauss
parents: 22422
diff changeset
  2161
  "(x,y)\<in> set (zip xs ys) \<Longrightarrow> x \<in> set xs"
db930e490fe5 added another rule for simultaneous induction, and lemmas for zip
krauss
parents: 22422
diff changeset
  2162
by (induct xs ys rule:list_induct2') auto
db930e490fe5 added another rule for simultaneous induction, and lemmas for zip
krauss
parents: 22422
diff changeset
  2163
db930e490fe5 added another rule for simultaneous induction, and lemmas for zip
krauss
parents: 22422
diff changeset
  2164
lemma set_zip_rightD:
db930e490fe5 added another rule for simultaneous induction, and lemmas for zip
krauss
parents: 22422
diff changeset
  2165
  "(x,y)\<in> set (zip xs ys) \<Longrightarrow> y \<in> set ys"
db930e490fe5 added another rule for simultaneous induction, and lemmas for zip
krauss
parents: 22422
diff changeset
  2166
by (induct xs ys rule:list_induct2') auto
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  2167
23983
79dc793bec43 Added lemmas
nipkow
parents: 23971
diff changeset
  2168
lemma in_set_zipE:
79dc793bec43 Added lemmas
nipkow
parents: 23971
diff changeset
  2169
  "(x,y) : set(zip xs ys) \<Longrightarrow> (\<lbrakk> x : set xs; y : set ys \<rbrakk> \<Longrightarrow> R) \<Longrightarrow> R"
79dc793bec43 Added lemmas
nipkow
parents: 23971
diff changeset
  2170
by(blast dest: set_zip_leftD set_zip_rightD)
79dc793bec43 Added lemmas
nipkow
parents: 23971
diff changeset
  2171
29829
9acb915a62fa code theorems for nth, list_update
haftmann
parents: 29827
diff changeset
  2172
lemma zip_map_fst_snd:
9acb915a62fa code theorems for nth, list_update
haftmann
parents: 29827
diff changeset
  2173
  "zip (map fst zs) (map snd zs) = zs"
9acb915a62fa code theorems for nth, list_update
haftmann
parents: 29827
diff changeset
  2174
  by (induct zs) simp_all
9acb915a62fa code theorems for nth, list_update
haftmann
parents: 29827
diff changeset
  2175
9acb915a62fa code theorems for nth, list_update
haftmann
parents: 29827
diff changeset
  2176
lemma zip_eq_conv:
9acb915a62fa code theorems for nth, list_update
haftmann
parents: 29827
diff changeset
  2177
  "length xs = length ys \<Longrightarrow> zip xs ys = zs \<longleftrightarrow> map fst zs = xs \<and> map snd zs = ys"
9acb915a62fa code theorems for nth, list_update
haftmann
parents: 29827
diff changeset
  2178
  by (auto simp add: zip_map_fst_snd)
9acb915a62fa code theorems for nth, list_update
haftmann
parents: 29827
diff changeset
  2179
35115
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
  2180
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15307
diff changeset
  2181
subsubsection {* @{text list_all2} *}
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2182
14316
91b897b9a2dc added some [intro?] and [trans] for list_all2 lemmas
kleing
parents: 14302
diff changeset
  2183
lemma list_all2_lengthD [intro?]: 
91b897b9a2dc added some [intro?] and [trans] for list_all2 lemmas
kleing
parents: 14302
diff changeset
  2184
  "list_all2 P xs ys ==> length xs = length ys"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  2185
by (simp add: list_all2_def)
19607
07eeb832f28d introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents: 19585
diff changeset
  2186
19787
b949911ecff5 improved code lemmas
haftmann
parents: 19770
diff changeset
  2187
lemma list_all2_Nil [iff, code]: "list_all2 P [] ys = (ys = [])"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  2188
by (simp add: list_all2_def)
19607
07eeb832f28d introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents: 19585
diff changeset
  2189
19787
b949911ecff5 improved code lemmas
haftmann
parents: 19770
diff changeset
  2190
lemma list_all2_Nil2 [iff, code]: "list_all2 P xs [] = (xs = [])"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  2191
by (simp add: list_all2_def)
19607
07eeb832f28d introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents: 19585
diff changeset
  2192
07eeb832f28d introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents: 19585
diff changeset
  2193
lemma list_all2_Cons [iff, code]:
07eeb832f28d introduced characters for code generator; some improved code lemmas for some list functions
haftmann
parents: 19585
diff changeset
  2194
  "list_all2 P (x # xs) (y # ys) = (P x y \<and> list_all2 P xs ys)"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  2195
by (auto simp add: list_all2_def)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2196
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2197
lemma list_all2_Cons1:
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2198
"list_all2 P (x # xs) ys = (\<exists>z zs. ys = z # zs \<and> P x z \<and> list_all2 P xs zs)"
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2199
by (cases ys) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2200
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2201
lemma list_all2_Cons2:
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2202
"list_all2 P xs (y # ys) = (\<exists>z zs. xs = z # zs \<and> P z y \<and> list_all2 P zs ys)"
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2203
by (cases xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2204
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  2205
lemma list_all2_rev [iff]:
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2206
"list_all2 P (rev xs) (rev ys) = list_all2 P xs ys"
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2207
by (simp add: list_all2_def zip_rev cong: conj_cong)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2208
13863
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2209
lemma list_all2_rev1:
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2210
"list_all2 P (rev xs) ys = list_all2 P xs (rev ys)"
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2211
by (subst list_all2_rev [symmetric]) simp
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2212
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2213
lemma list_all2_append1:
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2214
"list_all2 P (xs @ ys) zs =
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2215
(EX us vs. zs = us @ vs \<and> length us = length xs \<and> length vs = length ys \<and>
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2216
list_all2 P xs us \<and> list_all2 P ys vs)"
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2217
apply (simp add: list_all2_def zip_append1)
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2218
apply (rule iffI)
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2219
 apply (rule_tac x = "take (length xs) zs" in exI)
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2220
 apply (rule_tac x = "drop (length xs) zs" in exI)
14208
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
  2221
 apply (force split: nat_diff_split simp add: min_def, clarify)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2222
apply (simp add: ball_Un)
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2223
done
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2224
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2225
lemma list_all2_append2:
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2226
"list_all2 P xs (ys @ zs) =
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2227
(EX us vs. xs = us @ vs \<and> length us = length ys \<and> length vs = length zs \<and>
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2228
list_all2 P us ys \<and> list_all2 P vs zs)"
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2229
apply (simp add: list_all2_def zip_append2)
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2230
apply (rule iffI)
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2231
 apply (rule_tac x = "take (length ys) xs" in exI)
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2232
 apply (rule_tac x = "drop (length ys) xs" in exI)
14208
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
  2233
 apply (force split: nat_diff_split simp add: min_def, clarify)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2234
apply (simp add: ball_Un)
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2235
done
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2236
13863
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2237
lemma list_all2_append:
14247
cb32eb89bddd *** empty log message ***
nipkow
parents: 14208
diff changeset
  2238
  "length xs = length ys \<Longrightarrow>
cb32eb89bddd *** empty log message ***
nipkow
parents: 14208
diff changeset
  2239
  list_all2 P (xs@us) (ys@vs) = (list_all2 P xs ys \<and> list_all2 P us vs)"
cb32eb89bddd *** empty log message ***
nipkow
parents: 14208
diff changeset
  2240
by (induct rule:list_induct2, simp_all)
13863
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2241
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2242
lemma list_all2_appendI [intro?, trans]:
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2243
  "\<lbrakk> list_all2 P a b; list_all2 P c d \<rbrakk> \<Longrightarrow> list_all2 P (a@c) (b@d)"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  2244
by (simp add: list_all2_append list_all2_lengthD)
13863
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2245
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2246
lemma list_all2_conv_all_nth:
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2247
"list_all2 P xs ys =
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2248
(length xs = length ys \<and> (\<forall>i < length xs. P (xs!i) (ys!i)))"
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2249
by (force simp add: list_all2_def set_zip)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2250
13883
0451e0fb3f22 Re-structured some proofs in order to get rid of rule_format attribute.
berghofe
parents: 13863
diff changeset
  2251
lemma list_all2_trans:
0451e0fb3f22 Re-structured some proofs in order to get rid of rule_format attribute.
berghofe
parents: 13863
diff changeset
  2252
  assumes tr: "!!a b c. P1 a b ==> P2 b c ==> P3 a c"
0451e0fb3f22 Re-structured some proofs in order to get rid of rule_format attribute.
berghofe
parents: 13863
diff changeset
  2253
  shows "!!bs cs. list_all2 P1 as bs ==> list_all2 P2 bs cs ==> list_all2 P3 as cs"
0451e0fb3f22 Re-structured some proofs in order to get rid of rule_format attribute.
berghofe
parents: 13863
diff changeset
  2254
        (is "!!bs cs. PROP ?Q as bs cs")
0451e0fb3f22 Re-structured some proofs in order to get rid of rule_format attribute.
berghofe
parents: 13863
diff changeset
  2255
proof (induct as)
0451e0fb3f22 Re-structured some proofs in order to get rid of rule_format attribute.
berghofe
parents: 13863
diff changeset
  2256
  fix x xs bs assume I1: "!!bs cs. PROP ?Q xs bs cs"
0451e0fb3f22 Re-structured some proofs in order to get rid of rule_format attribute.
berghofe
parents: 13863
diff changeset
  2257
  show "!!cs. PROP ?Q (x # xs) bs cs"
0451e0fb3f22 Re-structured some proofs in order to get rid of rule_format attribute.
berghofe
parents: 13863
diff changeset
  2258
  proof (induct bs)
0451e0fb3f22 Re-structured some proofs in order to get rid of rule_format attribute.
berghofe
parents: 13863
diff changeset
  2259
    fix y ys cs assume I2: "!!cs. PROP ?Q (x # xs) ys cs"
0451e0fb3f22 Re-structured some proofs in order to get rid of rule_format attribute.
berghofe
parents: 13863
diff changeset
  2260
    show "PROP ?Q (x # xs) (y # ys) cs"
0451e0fb3f22 Re-structured some proofs in order to get rid of rule_format attribute.
berghofe
parents: 13863
diff changeset
  2261
      by (induct cs) (auto intro: tr I1 I2)
0451e0fb3f22 Re-structured some proofs in order to get rid of rule_format attribute.
berghofe
parents: 13863
diff changeset
  2262
  qed simp
0451e0fb3f22 Re-structured some proofs in order to get rid of rule_format attribute.
berghofe
parents: 13863
diff changeset
  2263
qed simp
0451e0fb3f22 Re-structured some proofs in order to get rid of rule_format attribute.
berghofe
parents: 13863
diff changeset
  2264
13863
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2265
lemma list_all2_all_nthI [intro?]:
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2266
  "length a = length b \<Longrightarrow> (\<And>n. n < length a \<Longrightarrow> P (a!n) (b!n)) \<Longrightarrow> list_all2 P a b"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  2267
by (simp add: list_all2_conv_all_nth)
13863
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2268
14395
cc96cc06abf9 new theorem
paulson
parents: 14388
diff changeset
  2269
lemma list_all2I:
cc96cc06abf9 new theorem
paulson
parents: 14388
diff changeset
  2270
  "\<forall>x \<in> set (zip a b). split P x \<Longrightarrow> length a = length b \<Longrightarrow> list_all2 P a b"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  2271
by (simp add: list_all2_def)
14395
cc96cc06abf9 new theorem
paulson
parents: 14388
diff changeset
  2272
14328
fd063037fdf5 list_all2_nthD no good as [intro?]
kleing
parents: 14327
diff changeset
  2273
lemma list_all2_nthD:
13863
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2274
  "\<lbrakk> list_all2 P xs ys; p < size xs \<rbrakk> \<Longrightarrow> P (xs!p) (ys!p)"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  2275
by (simp add: list_all2_conv_all_nth)
13863
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2276
14302
6c24235e8d5d *** empty log message ***
nipkow
parents: 14300
diff changeset
  2277
lemma list_all2_nthD2:
6c24235e8d5d *** empty log message ***
nipkow
parents: 14300
diff changeset
  2278
  "\<lbrakk>list_all2 P xs ys; p < size ys\<rbrakk> \<Longrightarrow> P (xs!p) (ys!p)"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  2279
by (frule list_all2_lengthD) (auto intro: list_all2_nthD)
14302
6c24235e8d5d *** empty log message ***
nipkow
parents: 14300
diff changeset
  2280
13863
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2281
lemma list_all2_map1: 
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2282
  "list_all2 P (map f as) bs = list_all2 (\<lambda>x y. P (f x) y) as bs"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  2283
by (simp add: list_all2_conv_all_nth)
13863
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2284
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2285
lemma list_all2_map2: 
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2286
  "list_all2 P as (map f bs) = list_all2 (\<lambda>x y. P x (f y)) as bs"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  2287
by (auto simp add: list_all2_conv_all_nth)
13863
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2288
14316
91b897b9a2dc added some [intro?] and [trans] for list_all2 lemmas
kleing
parents: 14302
diff changeset
  2289
lemma list_all2_refl [intro?]:
13863
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2290
  "(\<And>x. P x x) \<Longrightarrow> list_all2 P xs xs"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  2291
by (simp add: list_all2_conv_all_nth)
13863
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2292
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2293
lemma list_all2_update_cong:
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2294
  "\<lbrakk> i<size xs; list_all2 P xs ys; P x y \<rbrakk> \<Longrightarrow> list_all2 P (xs[i:=x]) (ys[i:=y])"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  2295
by (simp add: list_all2_conv_all_nth nth_list_update)
13863
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2296
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2297
lemma list_all2_update_cong2:
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2298
  "\<lbrakk>list_all2 P xs ys; P x y; i < length ys\<rbrakk> \<Longrightarrow> list_all2 P (xs[i:=x]) (ys[i:=y])"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  2299
by (simp add: list_all2_lengthD list_all2_update_cong)
13863
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2300
14302
6c24235e8d5d *** empty log message ***
nipkow
parents: 14300
diff changeset
  2301
lemma list_all2_takeI [simp,intro?]:
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2302
  "list_all2 P xs ys \<Longrightarrow> list_all2 P (take n xs) (take n ys)"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2303
apply (induct xs arbitrary: n ys)
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2304
 apply simp
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2305
apply (clarsimp simp add: list_all2_Cons1)
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2306
apply (case_tac n)
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2307
apply auto
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2308
done
14302
6c24235e8d5d *** empty log message ***
nipkow
parents: 14300
diff changeset
  2309
6c24235e8d5d *** empty log message ***
nipkow
parents: 14300
diff changeset
  2310
lemma list_all2_dropI [simp,intro?]:
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2311
  "list_all2 P as bs \<Longrightarrow> list_all2 P (drop n as) (drop n bs)"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2312
apply (induct as arbitrary: n bs, simp)
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2313
apply (clarsimp simp add: list_all2_Cons1)
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2314
apply (case_tac n, simp, simp)
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2315
done
13863
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2316
14327
9cd4dea593e3 list_all2_mono should not be [trans]
kleing
parents: 14316
diff changeset
  2317
lemma list_all2_mono [intro?]:
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2318
  "list_all2 P xs ys \<Longrightarrow> (\<And>xs ys. P xs ys \<Longrightarrow> Q xs ys) \<Longrightarrow> list_all2 Q xs ys"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2319
apply (induct xs arbitrary: ys, simp)
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2320
apply (case_tac ys, auto)
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2321
done
13863
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2322
22551
e52f5400e331 paraphrasing equality
haftmann
parents: 22539
diff changeset
  2323
lemma list_all2_eq:
e52f5400e331 paraphrasing equality
haftmann
parents: 22539
diff changeset
  2324
  "xs = ys \<longleftrightarrow> list_all2 (op =) xs ys"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  2325
by (induct xs ys rule: list_induct2') auto
22551
e52f5400e331 paraphrasing equality
haftmann
parents: 22539
diff changeset
  2326
40230
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  2327
lemma list_eq_iff_zip_eq:
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  2328
  "xs = ys \<longleftrightarrow> length xs = length ys \<and> (\<forall>(x,y) \<in> set (zip xs ys). x = y)"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  2329
by(auto simp add: set_zip list_all2_eq list_all2_conv_all_nth cong: conj_cong)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  2330
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  2331
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15307
diff changeset
  2332
subsubsection {* @{text foldl} and @{text foldr} *}
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  2333
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  2334
lemma foldl_append [simp]:
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2335
  "foldl f a (xs @ ys) = foldl f (foldl f a xs) ys"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2336
by (induct xs arbitrary: a) auto
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  2337
14402
4201e1916482 moved lemmas from MicroJava/Comp/AuxLemmas.thy to List.thy
nipkow
parents: 14395
diff changeset
  2338
lemma foldr_append[simp]: "foldr f (xs @ ys) a = foldr f xs (foldr f ys a)"
4201e1916482 moved lemmas from MicroJava/Comp/AuxLemmas.thy to List.thy
nipkow
parents: 14395
diff changeset
  2339
by (induct xs) auto
4201e1916482 moved lemmas from MicroJava/Comp/AuxLemmas.thy to List.thy
nipkow
parents: 14395
diff changeset
  2340
23096
423ad2fe9f76 *** empty log message ***
nipkow
parents: 23060
diff changeset
  2341
lemma foldr_map: "foldr g (map f xs) a = foldr (g o f) xs a"
423ad2fe9f76 *** empty log message ***
nipkow
parents: 23060
diff changeset
  2342
by(induct xs) simp_all
423ad2fe9f76 *** empty log message ***
nipkow
parents: 23060
diff changeset
  2343
24449
2f05cb7fed85 added (code) lemmas for setsum and foldl
nipkow
parents: 24349
diff changeset
  2344
text{* For efficient code generation: avoid intermediate list. *}
31998
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31930
diff changeset
  2345
lemma foldl_map[code_unfold]:
24449
2f05cb7fed85 added (code) lemmas for setsum and foldl
nipkow
parents: 24349
diff changeset
  2346
  "foldl g a (map f xs) = foldl (%a x. g a (f x)) a xs"
23096
423ad2fe9f76 *** empty log message ***
nipkow
parents: 23060
diff changeset
  2347
by(induct xs arbitrary:a) simp_all
423ad2fe9f76 *** empty log message ***
nipkow
parents: 23060
diff changeset
  2348
34978
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  2349
lemma foldl_apply:
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  2350
  assumes "\<And>x. x \<in> set xs \<Longrightarrow> f x \<circ> h = h \<circ> g x"
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  2351
  shows "foldl (\<lambda>s x. f x s) (h s) xs = h (foldl (\<lambda>s x. g x s) s xs)"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39272
diff changeset
  2352
  by (rule sym, insert assms, induct xs arbitrary: s) (simp_all add: fun_eq_iff)
31930
3107b9af1fb3 lemma foldl_apply_inv
haftmann
parents: 31784
diff changeset
  2353
44013
5cfc1c36ae97 moved recdef package to HOL/Library/Old_Recdef.thy
krauss
parents: 43594
diff changeset
  2354
lemma foldl_cong [fundef_cong]:
18336
1a2e30b37ed3 Added recdef congruence rules for bounded quantifiers and commonly used
krauss
parents: 18049
diff changeset
  2355
  "[| a = b; l = k; !!a x. x : set l ==> f a x = g a x |] 
1a2e30b37ed3 Added recdef congruence rules for bounded quantifiers and commonly used
krauss
parents: 18049
diff changeset
  2356
  ==> foldl f a l = foldl g b k"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  2357
by (induct k arbitrary: a b l) simp_all
18336
1a2e30b37ed3 Added recdef congruence rules for bounded quantifiers and commonly used
krauss
parents: 18049
diff changeset
  2358
44013
5cfc1c36ae97 moved recdef package to HOL/Library/Old_Recdef.thy
krauss
parents: 43594
diff changeset
  2359
lemma foldr_cong [fundef_cong]:
18336
1a2e30b37ed3 Added recdef congruence rules for bounded quantifiers and commonly used
krauss
parents: 18049
diff changeset
  2360
  "[| a = b; l = k; !!a x. x : set l ==> f x a = g x a |] 
1a2e30b37ed3 Added recdef congruence rules for bounded quantifiers and commonly used
krauss
parents: 18049
diff changeset
  2361
  ==> foldr f l a = foldr g k b"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  2362
by (induct k arbitrary: a b l) simp_all
18336
1a2e30b37ed3 Added recdef congruence rules for bounded quantifiers and commonly used
krauss
parents: 18049
diff changeset
  2363
35195
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  2364
lemma foldl_fun_comm:
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  2365
  assumes "\<And>x y s. f (f s x) y = f (f s y) x"
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  2366
  shows "f (foldl f s xs) x = foldl f (f s x) xs"
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  2367
  by (induct xs arbitrary: s)
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  2368
    (simp_all add: assms)
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  2369
24449
2f05cb7fed85 added (code) lemmas for setsum and foldl
nipkow
parents: 24349
diff changeset
  2370
lemma (in semigroup_add) foldl_assoc:
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24902
diff changeset
  2371
shows "foldl op+ (x+y) zs = x + (foldl op+ y zs)"
24449
2f05cb7fed85 added (code) lemmas for setsum and foldl
nipkow
parents: 24349
diff changeset
  2372
by (induct zs arbitrary: y) (simp_all add:add_assoc)
2f05cb7fed85 added (code) lemmas for setsum and foldl
nipkow
parents: 24349
diff changeset
  2373
2f05cb7fed85 added (code) lemmas for setsum and foldl
nipkow
parents: 24349
diff changeset
  2374
lemma (in monoid_add) foldl_absorb0:
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24902
diff changeset
  2375
shows "x + (foldl op+ 0 zs) = foldl op+ x zs"
24449
2f05cb7fed85 added (code) lemmas for setsum and foldl
nipkow
parents: 24349
diff changeset
  2376
by (induct zs) (simp_all add:foldl_assoc)
2f05cb7fed85 added (code) lemmas for setsum and foldl
nipkow
parents: 24349
diff changeset
  2377
35195
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  2378
lemma foldl_rev:
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  2379
  assumes "\<And>x y s. f (f s x) y = f (f s y) x"
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  2380
  shows "foldl f s (rev xs) = foldl f s xs"
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  2381
proof (induct xs arbitrary: s)
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  2382
  case Nil then show ?case by simp
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  2383
next
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  2384
  case (Cons x xs) with assms show ?case by (simp add: foldl_fun_comm)
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  2385
qed
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  2386
37605
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  2387
lemma rev_foldl_cons [code]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  2388
  "rev xs = foldl (\<lambda>xs x. x # xs) [] xs"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  2389
proof (induct xs)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  2390
  case Nil then show ?case by simp
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  2391
next
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  2392
  case Cons
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  2393
  {
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  2394
    fix x xs ys
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  2395
    have "foldl (\<lambda>xs x. x # xs) ys xs @ [x]
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  2396
      = foldl (\<lambda>xs x. x # xs) (ys @ [x]) xs"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  2397
    by (induct xs arbitrary: ys) auto
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  2398
  }
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  2399
  note aux = this
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  2400
  show ?case by (induct xs) (auto simp add: Cons aux)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  2401
qed
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  2402
24449
2f05cb7fed85 added (code) lemmas for setsum and foldl
nipkow
parents: 24349
diff changeset
  2403
39774
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2404
text{* The ``Third Duality Theorem'' in Bird \& Wadler: *}
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2405
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2406
lemma foldr_foldl:
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2407
  "foldr f xs a = foldl (%x y. f y x) a (rev xs)"
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2408
  by (induct xs) auto
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2409
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2410
lemma foldl_foldr:
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2411
  "foldl f a xs = foldr (%x y. f y x) (rev xs) a"
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2412
  by (simp add: foldr_foldl [of "%x y. f y x" "rev xs"])
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2413
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2414
23096
423ad2fe9f76 *** empty log message ***
nipkow
parents: 23060
diff changeset
  2415
text{* The ``First Duality Theorem'' in Bird \& Wadler: *}
423ad2fe9f76 *** empty log message ***
nipkow
parents: 23060
diff changeset
  2416
39774
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2417
lemma (in monoid_add) foldl_foldr1_lemma:
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2418
  "foldl op + a xs = a + foldr op + xs 0"
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2419
  by (induct xs arbitrary: a) (auto simp: add_assoc)
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2420
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2421
corollary (in monoid_add) foldl_foldr1:
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2422
  "foldl op + 0 xs = foldr op + xs 0"
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2423
  by (simp add: foldl_foldr1_lemma)
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2424
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2425
lemma (in ab_semigroup_add) foldr_conv_foldl:
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2426
  "foldr op + xs a = foldl op + a xs"
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2427
  by (induct xs) (simp_all add: foldl_assoc add.commute)
24471
d7cf53c1085f removed unused theorems ; added lifting properties for foldr and foldl
chaieb
parents: 24461
diff changeset
  2428
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  2429
text {*
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2430
Note: @{text "n \<le> foldl (op +) n ns"} looks simpler, but is more
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2431
difficult to use because it requires an additional transitivity step.
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  2432
*}
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  2433
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2434
lemma start_le_sum: "(m::nat) <= n ==> m <= foldl (op +) n ns"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2435
by (induct ns arbitrary: n) auto
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2436
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2437
lemma elem_le_sum: "(n::nat) : set ns ==> n <= foldl (op +) 0 ns"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2438
by (force intro: start_le_sum simp add: in_set_conv_decomp)
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  2439
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  2440
lemma sum_eq_0_conv [iff]:
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2441
  "(foldl (op +) (m::nat) ns = 0) = (m = 0 \<and> (\<forall>n \<in> set ns. n = 0))"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2442
by (induct ns arbitrary: m) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2443
24471
d7cf53c1085f removed unused theorems ; added lifting properties for foldr and foldl
chaieb
parents: 24461
diff changeset
  2444
lemma foldr_invariant: 
d7cf53c1085f removed unused theorems ; added lifting properties for foldr and foldl
chaieb
parents: 24461
diff changeset
  2445
  "\<lbrakk>Q x ; \<forall> x\<in> set xs. P x; \<forall> x y. P x \<and> Q y \<longrightarrow> Q (f x y) \<rbrakk> \<Longrightarrow> Q (foldr f xs x)"
d7cf53c1085f removed unused theorems ; added lifting properties for foldr and foldl
chaieb
parents: 24461
diff changeset
  2446
  by (induct xs, simp_all)
d7cf53c1085f removed unused theorems ; added lifting properties for foldr and foldl
chaieb
parents: 24461
diff changeset
  2447
d7cf53c1085f removed unused theorems ; added lifting properties for foldr and foldl
chaieb
parents: 24461
diff changeset
  2448
lemma foldl_invariant: 
d7cf53c1085f removed unused theorems ; added lifting properties for foldr and foldl
chaieb
parents: 24461
diff changeset
  2449
  "\<lbrakk>Q x ; \<forall> x\<in> set xs. P x; \<forall> x y. P x \<and> Q y \<longrightarrow> Q (f y x) \<rbrakk> \<Longrightarrow> Q (foldl f x xs)"
d7cf53c1085f removed unused theorems ; added lifting properties for foldr and foldl
chaieb
parents: 24461
diff changeset
  2450
  by (induct xs arbitrary: x, simp_all)
d7cf53c1085f removed unused theorems ; added lifting properties for foldr and foldl
chaieb
parents: 24461
diff changeset
  2451
34978
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  2452
lemma foldl_weak_invariant:
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  2453
  assumes "P s"
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  2454
    and "\<And>s x. x \<in> set xs \<Longrightarrow> P s \<Longrightarrow> P (f s x)"
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  2455
  shows "P (foldl f s xs)"
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  2456
  using assms by (induct xs arbitrary: s) simp_all
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  2457
31455
2754a0dadccc lemma about List.foldl and Finite_Set.fold
haftmann
parents: 31363
diff changeset
  2458
text {* @{const foldl} and @{const concat} *}
24449
2f05cb7fed85 added (code) lemmas for setsum and foldl
nipkow
parents: 24349
diff changeset
  2459
2f05cb7fed85 added (code) lemmas for setsum and foldl
nipkow
parents: 24349
diff changeset
  2460
lemma foldl_conv_concat:
29782
02e76245e5af dropped global Nil/Append interpretation
haftmann
parents: 29626
diff changeset
  2461
  "foldl (op @) xs xss = xs @ concat xss"
02e76245e5af dropped global Nil/Append interpretation
haftmann
parents: 29626
diff changeset
  2462
proof (induct xss arbitrary: xs)
02e76245e5af dropped global Nil/Append interpretation
haftmann
parents: 29626
diff changeset
  2463
  case Nil show ?case by simp
02e76245e5af dropped global Nil/Append interpretation
haftmann
parents: 29626
diff changeset
  2464
next
35267
8dfd816713c6 moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents: 35217
diff changeset
  2465
  interpret monoid_add "op @" "[]" proof qed simp_all
29782
02e76245e5af dropped global Nil/Append interpretation
haftmann
parents: 29626
diff changeset
  2466
  case Cons then show ?case by (simp add: foldl_absorb0)
02e76245e5af dropped global Nil/Append interpretation
haftmann
parents: 29626
diff changeset
  2467
qed
02e76245e5af dropped global Nil/Append interpretation
haftmann
parents: 29626
diff changeset
  2468
02e76245e5af dropped global Nil/Append interpretation
haftmann
parents: 29626
diff changeset
  2469
lemma concat_conv_foldl: "concat xss = foldl (op @) [] xss"
02e76245e5af dropped global Nil/Append interpretation
haftmann
parents: 29626
diff changeset
  2470
  by (simp add: foldl_conv_concat)
02e76245e5af dropped global Nil/Append interpretation
haftmann
parents: 29626
diff changeset
  2471
31455
2754a0dadccc lemma about List.foldl and Finite_Set.fold
haftmann
parents: 31363
diff changeset
  2472
text {* @{const Finite_Set.fold} and @{const foldl} *}
2754a0dadccc lemma about List.foldl and Finite_Set.fold
haftmann
parents: 31363
diff changeset
  2473
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42809
diff changeset
  2474
lemma (in comp_fun_commute) fold_set_remdups:
35195
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  2475
  "fold f y (set xs) = foldl (\<lambda>y x. f x y) y (remdups xs)"
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  2476
  by (rule sym, induct xs arbitrary: y) (simp_all add: fold_fun_comm insert_absorb)
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  2477
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42809
diff changeset
  2478
lemma (in comp_fun_idem) fold_set:
31455
2754a0dadccc lemma about List.foldl and Finite_Set.fold
haftmann
parents: 31363
diff changeset
  2479
  "fold f y (set xs) = foldl (\<lambda>y x. f x y) y xs"
2754a0dadccc lemma about List.foldl and Finite_Set.fold
haftmann
parents: 31363
diff changeset
  2480
  by (rule sym, induct xs arbitrary: y) (simp_all add: fold_fun_comm)
2754a0dadccc lemma about List.foldl and Finite_Set.fold
haftmann
parents: 31363
diff changeset
  2481
32681
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2482
lemma (in ab_semigroup_idem_mult) fold1_set:
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2483
  assumes "xs \<noteq> []"
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2484
  shows "fold1 times (set xs) = foldl times (hd xs) (tl xs)"
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2485
proof -
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42809
diff changeset
  2486
  interpret comp_fun_idem times by (fact comp_fun_idem)
32681
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2487
  from assms obtain y ys where xs: "xs = y # ys"
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2488
    by (cases xs) auto
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2489
  show ?thesis
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2490
  proof (cases "set ys = {}")
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2491
    case True with xs show ?thesis by simp
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2492
  next
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2493
    case False
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2494
    then have "fold1 times (insert y (set ys)) = fold times y (set ys)"
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2495
      by (simp only: finite_set fold1_eq_fold_idem)
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2496
    with xs show ?thesis by (simp add: fold_set mult_commute)
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2497
  qed
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2498
qed
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2499
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2500
lemma (in lattice) Inf_fin_set_fold [code_unfold]:
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2501
  "Inf_fin (set (x # xs)) = foldl inf x xs"
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2502
proof -
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2503
  interpret ab_semigroup_idem_mult "inf :: 'a \<Rightarrow> 'a \<Rightarrow> 'a"
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2504
    by (fact ab_semigroup_idem_mult_inf)
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2505
  show ?thesis
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2506
    by (simp add: Inf_fin_def fold1_set del: set.simps)
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2507
qed
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2508
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2509
lemma (in lattice) Sup_fin_set_fold [code_unfold]:
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2510
  "Sup_fin (set (x # xs)) = foldl sup x xs"
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2511
proof -
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2512
  interpret ab_semigroup_idem_mult "sup :: 'a \<Rightarrow> 'a \<Rightarrow> 'a"
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2513
    by (fact ab_semigroup_idem_mult_sup)
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2514
  show ?thesis
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2515
    by (simp add: Sup_fin_def fold1_set del: set.simps)
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2516
qed
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2517
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2518
lemma (in linorder) Min_fin_set_fold [code_unfold]:
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2519
  "Min (set (x # xs)) = foldl min x xs"
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2520
proof -
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2521
  interpret ab_semigroup_idem_mult "min :: 'a \<Rightarrow> 'a \<Rightarrow> 'a"
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2522
    by (fact ab_semigroup_idem_mult_min)
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2523
  show ?thesis
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2524
    by (simp add: Min_def fold1_set del: set.simps)
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2525
qed
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2526
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2527
lemma (in linorder) Max_fin_set_fold [code_unfold]:
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2528
  "Max (set (x # xs)) = foldl max x xs"
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2529
proof -
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2530
  interpret ab_semigroup_idem_mult "max :: 'a \<Rightarrow> 'a \<Rightarrow> 'a"
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2531
    by (fact ab_semigroup_idem_mult_max)
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2532
  show ?thesis
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2533
    by (simp add: Max_def fold1_set del: set.simps)
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2534
qed
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2535
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2536
lemma (in complete_lattice) Inf_set_fold [code_unfold]:
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2537
  "Inf (set xs) = foldl inf top xs"
34007
aea892559fc5 tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents: 33972
diff changeset
  2538
proof -
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42809
diff changeset
  2539
  interpret comp_fun_idem "inf :: 'a \<Rightarrow> 'a \<Rightarrow> 'a"
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42809
diff changeset
  2540
    by (fact comp_fun_idem_inf)
34007
aea892559fc5 tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents: 33972
diff changeset
  2541
  show ?thesis by (simp add: Inf_fold_inf fold_set inf_commute)
aea892559fc5 tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents: 33972
diff changeset
  2542
qed
32681
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2543
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2544
lemma (in complete_lattice) Sup_set_fold [code_unfold]:
adeac3cbb659 lemma relating fold1 and foldl; code_unfold rules for Inf_fin, Sup_fin, Min, Max, Inf, Sup
haftmann
parents: 32422
diff changeset
  2545
  "Sup (set xs) = foldl sup bot xs"
34007
aea892559fc5 tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents: 33972
diff changeset
  2546
proof -
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42809
diff changeset
  2547
  interpret comp_fun_idem "sup :: 'a \<Rightarrow> 'a \<Rightarrow> 'a"
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42809
diff changeset
  2548
    by (fact comp_fun_idem_sup)
34007
aea892559fc5 tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents: 33972
diff changeset
  2549
  show ?thesis by (simp add: Sup_fold_sup fold_set sup_commute)
aea892559fc5 tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents: 33972
diff changeset
  2550
qed
aea892559fc5 tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents: 33972
diff changeset
  2551
aea892559fc5 tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents: 33972
diff changeset
  2552
lemma (in complete_lattice) INFI_set_fold:
aea892559fc5 tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents: 33972
diff changeset
  2553
  "INFI (set xs) f = foldl (\<lambda>y x. inf (f x) y) top xs"
44928
7ef6505bde7f renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents: 44921
diff changeset
  2554
  unfolding INF_def set_map [symmetric] Inf_set_fold foldl_map
34007
aea892559fc5 tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents: 33972
diff changeset
  2555
    by (simp add: inf_commute)
aea892559fc5 tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents: 33972
diff changeset
  2556
aea892559fc5 tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents: 33972
diff changeset
  2557
lemma (in complete_lattice) SUPR_set_fold:
aea892559fc5 tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents: 33972
diff changeset
  2558
  "SUPR (set xs) f = foldl (\<lambda>y x. sup (f x) y) bot xs"
44928
7ef6505bde7f renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents: 44921
diff changeset
  2559
  unfolding SUP_def set_map [symmetric] Sup_set_fold foldl_map
34007
aea892559fc5 tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents: 33972
diff changeset
  2560
    by (simp add: sup_commute)
31455
2754a0dadccc lemma about List.foldl and Finite_Set.fold
haftmann
parents: 31363
diff changeset
  2561
35115
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
  2562
24645
1af302128da2 Generalized [_.._] from nat to linear orders
nipkow
parents: 24640
diff changeset
  2563
subsubsection {* @{text upt} *}
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2564
17090
603f23d71ada small mods to code lemmas
nipkow
parents: 17086
diff changeset
  2565
lemma upt_rec[code]: "[i..<j] = (if i<j then i#[Suc i..<j] else [])"
603f23d71ada small mods to code lemmas
nipkow
parents: 17086
diff changeset
  2566
-- {* simp does not terminate! *}
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2567
by (induct j) auto
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  2568
45607
16b4f5774621 eliminated obsolete "standard";
wenzelm
parents: 45181
diff changeset
  2569
lemmas upt_rec_number_of[simp] = upt_rec[of "number_of m" "number_of n"] for m n
32005
c369417be055 made upt/upto executable on nat/int by simp
nipkow
parents: 31998
diff changeset
  2570
15425
6356d2523f73 [ .. (] -> [ ..< ]
nipkow
parents: 15392
diff changeset
  2571
lemma upt_conv_Nil [simp]: "j <= i ==> [i..<j] = []"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2572
by (subst upt_rec) simp
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2573
15425
6356d2523f73 [ .. (] -> [ ..< ]
nipkow
parents: 15392
diff changeset
  2574
lemma upt_eq_Nil_conv[simp]: "([i..<j] = []) = (j = 0 \<or> j <= i)"
15281
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  2575
by(induct j)simp_all
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  2576
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  2577
lemma upt_eq_Cons_conv:
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2578
 "([i..<j] = x#xs) = (i < j & i = x & [i+1..<j] = xs)"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2579
apply(induct j arbitrary: x xs)
15281
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  2580
 apply simp
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  2581
apply(clarsimp simp add: append_eq_Cons_conv)
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  2582
apply arith
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  2583
done
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  2584
15425
6356d2523f73 [ .. (] -> [ ..< ]
nipkow
parents: 15392
diff changeset
  2585
lemma upt_Suc_append: "i <= j ==> [i..<(Suc j)] = [i..<j]@[j]"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2586
-- {* Only needed if @{text upt_Suc} is deleted from the simpset. *}
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2587
by simp
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2588
15425
6356d2523f73 [ .. (] -> [ ..< ]
nipkow
parents: 15392
diff changeset
  2589
lemma upt_conv_Cons: "i < j ==> [i..<j] = i # [Suc i..<j]"
26734
a92057c1ee21 dropped some metis calls
haftmann
parents: 26584
diff changeset
  2590
  by (simp add: upt_rec)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2591
15425
6356d2523f73 [ .. (] -> [ ..< ]
nipkow
parents: 15392
diff changeset
  2592
lemma upt_add_eq_append: "i<=j ==> [i..<j+k] = [i..<j]@[j..<j+k]"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2593
-- {* LOOPS as a simprule, since @{text "j <= j"}. *}
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2594
by (induct k) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2595
15425
6356d2523f73 [ .. (] -> [ ..< ]
nipkow
parents: 15392
diff changeset
  2596
lemma length_upt [simp]: "length [i..<j] = j - i"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2597
by (induct j) (auto simp add: Suc_diff_le)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2598
15425
6356d2523f73 [ .. (] -> [ ..< ]
nipkow
parents: 15392
diff changeset
  2599
lemma nth_upt [simp]: "i + k < j ==> [i..<j] ! k = i + k"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2600
apply (induct j)
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2601
apply (auto simp add: less_Suc_eq nth_append split: nat_diff_split)
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2602
done
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2603
17906
719364f5179b added 2 lemmas
nipkow
parents: 17877
diff changeset
  2604
719364f5179b added 2 lemmas
nipkow
parents: 17877
diff changeset
  2605
lemma hd_upt[simp]: "i < j \<Longrightarrow> hd[i..<j] = i"
719364f5179b added 2 lemmas
nipkow
parents: 17877
diff changeset
  2606
by(simp add:upt_conv_Cons)
719364f5179b added 2 lemmas
nipkow
parents: 17877
diff changeset
  2607
719364f5179b added 2 lemmas
nipkow
parents: 17877
diff changeset
  2608
lemma last_upt[simp]: "i < j \<Longrightarrow> last[i..<j] = j - 1"
719364f5179b added 2 lemmas
nipkow
parents: 17877
diff changeset
  2609
apply(cases j)
719364f5179b added 2 lemmas
nipkow
parents: 17877
diff changeset
  2610
 apply simp
719364f5179b added 2 lemmas
nipkow
parents: 17877
diff changeset
  2611
by(simp add:upt_Suc_append)
719364f5179b added 2 lemmas
nipkow
parents: 17877
diff changeset
  2612
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2613
lemma take_upt [simp]: "i+m <= n ==> take m [i..<n] = [i..<i+m]"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2614
apply (induct m arbitrary: i, simp)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2615
apply (subst upt_rec)
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2616
apply (rule sym)
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2617
apply (subst upt_rec)
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2618
apply (simp del: upt.simps)
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2619
done
3507
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
  2620
17501
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2621
lemma drop_upt[simp]: "drop m [i..<j] = [i+m..<j]"
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2622
apply(induct j)
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2623
apply auto
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2624
done
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2625
24645
1af302128da2 Generalized [_.._] from nat to linear orders
nipkow
parents: 24640
diff changeset
  2626
lemma map_Suc_upt: "map Suc [m..<n] = [Suc m..<Suc n]"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2627
by (induct n) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2628
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2629
lemma nth_map_upt: "i < n-m ==> (map f [m..<n]) ! i = f(m+i)"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2630
apply (induct n m  arbitrary: i rule: diff_induct)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2631
prefer 3 apply (subst map_Suc_upt[symmetric])
44921
58eef4843641 tuned proofs
huffman
parents: 44916
diff changeset
  2632
apply (auto simp add: less_diff_conv)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2633
done
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2634
13883
0451e0fb3f22 Re-structured some proofs in order to get rid of rule_format attribute.
berghofe
parents: 13863
diff changeset
  2635
lemma nth_take_lemma:
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2636
  "k <= length xs ==> k <= length ys ==>
13883
0451e0fb3f22 Re-structured some proofs in order to get rid of rule_format attribute.
berghofe
parents: 13863
diff changeset
  2637
     (!!i. i < k --> xs!i = ys!i) ==> take k xs = take k ys"
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2638
apply (atomize, induct k arbitrary: xs ys)
14208
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
  2639
apply (simp_all add: less_Suc_eq_0_disj all_conj_distrib, clarify)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2640
txt {* Both lists must be non-empty *}
14208
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
  2641
apply (case_tac xs, simp)
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
  2642
apply (case_tac ys, clarify)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2643
 apply (simp (no_asm_use))
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2644
apply clarify
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2645
txt {* prenexing's needed, not miniscoping *}
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2646
apply (simp (no_asm_use) add: all_simps [symmetric] del: all_simps)
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2647
apply blast
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2648
done
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2649
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2650
lemma nth_equalityI:
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2651
 "[| length xs = length ys; ALL i < length xs. xs!i = ys!i |] ==> xs = ys"
44921
58eef4843641 tuned proofs
huffman
parents: 44916
diff changeset
  2652
  by (frule nth_take_lemma [OF le_refl eq_imp_le]) simp_all
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  2653
24796
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  2654
lemma map_nth:
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  2655
  "map (\<lambda>i. xs ! i) [0..<length xs] = xs"
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  2656
  by (rule nth_equalityI, auto)
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  2657
13863
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2658
(* needs nth_equalityI *)
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2659
lemma list_all2_antisym:
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2660
  "\<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> 
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2661
  \<Longrightarrow> xs = ys"
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2662
  apply (simp add: list_all2_conv_all_nth) 
14208
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
  2663
  apply (rule nth_equalityI, blast, simp)
13863
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2664
  done
ec901a432ea1 more about list_all2
kleing
parents: 13737
diff changeset
  2665
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  2666
lemma take_equalityI: "(\<forall>i. take i xs = take i ys) ==> xs = ys"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2667
-- {* The famous take-lemma. *}
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2668
apply (drule_tac x = "max (length xs) (length ys)" in spec)
44921
58eef4843641 tuned proofs
huffman
parents: 44916
diff changeset
  2669
apply (simp add: le_max_iff_disj)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2670
done
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  2671
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  2672
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  2673
lemma take_Cons':
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  2674
     "take n (x # xs) = (if n = 0 then [] else x # take (n - 1) xs)"
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  2675
by (cases n) simp_all
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  2676
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  2677
lemma drop_Cons':
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  2678
     "drop n (x # xs) = (if n = 0 then x # xs else drop (n - 1) xs)"
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  2679
by (cases n) simp_all
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  2680
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  2681
lemma nth_Cons': "(x # xs)!n = (if n = 0 then x else xs!(n - 1))"
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  2682
by (cases n) simp_all
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  2683
45607
16b4f5774621 eliminated obsolete "standard";
wenzelm
parents: 45181
diff changeset
  2684
lemmas take_Cons_number_of = take_Cons'[of "number_of v"] for v
16b4f5774621 eliminated obsolete "standard";
wenzelm
parents: 45181
diff changeset
  2685
lemmas drop_Cons_number_of = drop_Cons'[of "number_of v"] for v
16b4f5774621 eliminated obsolete "standard";
wenzelm
parents: 45181
diff changeset
  2686
lemmas nth_Cons_number_of = nth_Cons'[of _ _ "number_of v"] for v
18622
4524643feecc theorems need names
paulson
parents: 18490
diff changeset
  2687
4524643feecc theorems need names
paulson
parents: 18490
diff changeset
  2688
declare take_Cons_number_of [simp] 
4524643feecc theorems need names
paulson
parents: 18490
diff changeset
  2689
        drop_Cons_number_of [simp] 
4524643feecc theorems need names
paulson
parents: 18490
diff changeset
  2690
        nth_Cons_number_of [simp] 
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  2691
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  2692
32415
1dddf2f64266 got rid of complicated class finite_intvl_succ and defined "upto" directly on int, the only instance of the class.
nipkow
parents: 32078
diff changeset
  2693
subsubsection {* @{text upto}: interval-list on @{typ int} *}
1dddf2f64266 got rid of complicated class finite_intvl_succ and defined "upto" directly on int, the only instance of the class.
nipkow
parents: 32078
diff changeset
  2694
1dddf2f64266 got rid of complicated class finite_intvl_succ and defined "upto" directly on int, the only instance of the class.
nipkow
parents: 32078
diff changeset
  2695
(* FIXME make upto tail recursive? *)
1dddf2f64266 got rid of complicated class finite_intvl_succ and defined "upto" directly on int, the only instance of the class.
nipkow
parents: 32078
diff changeset
  2696
1dddf2f64266 got rid of complicated class finite_intvl_succ and defined "upto" directly on int, the only instance of the class.
nipkow
parents: 32078
diff changeset
  2697
function upto :: "int \<Rightarrow> int \<Rightarrow> int list" ("(1[_../_])") where
1dddf2f64266 got rid of complicated class finite_intvl_succ and defined "upto" directly on int, the only instance of the class.
nipkow
parents: 32078
diff changeset
  2698
"upto i j = (if i \<le> j then i # [i+1..j] else [])"
1dddf2f64266 got rid of complicated class finite_intvl_succ and defined "upto" directly on int, the only instance of the class.
nipkow
parents: 32078
diff changeset
  2699
by auto
1dddf2f64266 got rid of complicated class finite_intvl_succ and defined "upto" directly on int, the only instance of the class.
nipkow
parents: 32078
diff changeset
  2700
termination
1dddf2f64266 got rid of complicated class finite_intvl_succ and defined "upto" directly on int, the only instance of the class.
nipkow
parents: 32078
diff changeset
  2701
by(relation "measure(%(i::int,j). nat(j - i + 1))") auto
1dddf2f64266 got rid of complicated class finite_intvl_succ and defined "upto" directly on int, the only instance of the class.
nipkow
parents: 32078
diff changeset
  2702
1dddf2f64266 got rid of complicated class finite_intvl_succ and defined "upto" directly on int, the only instance of the class.
nipkow
parents: 32078
diff changeset
  2703
declare upto.simps[code, simp del]
1dddf2f64266 got rid of complicated class finite_intvl_succ and defined "upto" directly on int, the only instance of the class.
nipkow
parents: 32078
diff changeset
  2704
45607
16b4f5774621 eliminated obsolete "standard";
wenzelm
parents: 45181
diff changeset
  2705
lemmas upto_rec_number_of[simp] = upto.simps[of "number_of m" "number_of n"] for m n
32415
1dddf2f64266 got rid of complicated class finite_intvl_succ and defined "upto" directly on int, the only instance of the class.
nipkow
parents: 32078
diff changeset
  2706
1dddf2f64266 got rid of complicated class finite_intvl_succ and defined "upto" directly on int, the only instance of the class.
nipkow
parents: 32078
diff changeset
  2707
lemma upto_empty[simp]: "j < i \<Longrightarrow> [i..j] = []"
1dddf2f64266 got rid of complicated class finite_intvl_succ and defined "upto" directly on int, the only instance of the class.
nipkow
parents: 32078
diff changeset
  2708
by(simp add: upto.simps)
1dddf2f64266 got rid of complicated class finite_intvl_succ and defined "upto" directly on int, the only instance of the class.
nipkow
parents: 32078
diff changeset
  2709
1dddf2f64266 got rid of complicated class finite_intvl_succ and defined "upto" directly on int, the only instance of the class.
nipkow
parents: 32078
diff changeset
  2710
lemma set_upto[simp]: "set[i..j] = {i..j}"
41463
edbf0a86fb1c adding simproc to rewrite list comprehensions to set comprehensions; adopting proofs
bulwahn
parents: 41372
diff changeset
  2711
proof(induct i j rule:upto.induct)
edbf0a86fb1c adding simproc to rewrite list comprehensions to set comprehensions; adopting proofs
bulwahn
parents: 41372
diff changeset
  2712
  case (1 i j)
edbf0a86fb1c adding simproc to rewrite list comprehensions to set comprehensions; adopting proofs
bulwahn
parents: 41372
diff changeset
  2713
  from this show ?case
edbf0a86fb1c adding simproc to rewrite list comprehensions to set comprehensions; adopting proofs
bulwahn
parents: 41372
diff changeset
  2714
    unfolding upto.simps[of i j] simp_from_to[of i j] by auto
edbf0a86fb1c adding simproc to rewrite list comprehensions to set comprehensions; adopting proofs
bulwahn
parents: 41372
diff changeset
  2715
qed
32415
1dddf2f64266 got rid of complicated class finite_intvl_succ and defined "upto" directly on int, the only instance of the class.
nipkow
parents: 32078
diff changeset
  2716
1dddf2f64266 got rid of complicated class finite_intvl_succ and defined "upto" directly on int, the only instance of the class.
nipkow
parents: 32078
diff changeset
  2717
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15307
diff changeset
  2718
subsubsection {* @{text "distinct"} and @{text remdups} *}
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  2719
40210
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  2720
lemma distinct_tl:
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  2721
  "distinct xs \<Longrightarrow> distinct (tl xs)"
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  2722
  by (cases xs) simp_all
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  2723
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  2724
lemma distinct_append [simp]:
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2725
"distinct (xs @ ys) = (distinct xs \<and> distinct ys \<and> set xs \<inter> set ys = {})"
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2726
by (induct xs) auto
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  2727
15305
0bd9eedaa301 added lemmas
nipkow
parents: 15304
diff changeset
  2728
lemma distinct_rev[simp]: "distinct(rev xs) = distinct xs"
0bd9eedaa301 added lemmas
nipkow
parents: 15304
diff changeset
  2729
by(induct xs) auto
0bd9eedaa301 added lemmas
nipkow
parents: 15304
diff changeset
  2730
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  2731
lemma set_remdups [simp]: "set (remdups xs) = set xs"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2732
by (induct xs) (auto simp add: insert_absorb)
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  2733
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  2734
lemma distinct_remdups [iff]: "distinct (remdups xs)"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2735
by (induct xs) auto
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  2736
25287
094dab519ff5 added lemmas
nipkow
parents: 25277
diff changeset
  2737
lemma distinct_remdups_id: "distinct xs ==> remdups xs = xs"
094dab519ff5 added lemmas
nipkow
parents: 25277
diff changeset
  2738
by (induct xs, auto)
094dab519ff5 added lemmas
nipkow
parents: 25277
diff changeset
  2739
26734
a92057c1ee21 dropped some metis calls
haftmann
parents: 26584
diff changeset
  2740
lemma remdups_id_iff_distinct [simp]: "remdups xs = xs \<longleftrightarrow> distinct xs"
a92057c1ee21 dropped some metis calls
haftmann
parents: 26584
diff changeset
  2741
by (metis distinct_remdups distinct_remdups_id)
25287
094dab519ff5 added lemmas
nipkow
parents: 25277
diff changeset
  2742
24566
2bfa0215904c added lemma
nipkow
parents: 24526
diff changeset
  2743
lemma finite_distinct_list: "finite A \<Longrightarrow> EX xs. set xs = A & distinct xs"
24632
779fc4fcbf8b metis now available in PreList
paulson
parents: 24617
diff changeset
  2744
by (metis distinct_remdups finite_list set_remdups)
24566
2bfa0215904c added lemma
nipkow
parents: 24526
diff changeset
  2745
15072
4861bf6af0b4 new material courtesy of Norbert Voelker
paulson
parents: 15064
diff changeset
  2746
lemma remdups_eq_nil_iff [simp]: "(remdups x = []) = (x = [])"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  2747
by (induct x, auto) 
15072
4861bf6af0b4 new material courtesy of Norbert Voelker
paulson
parents: 15064
diff changeset
  2748
4861bf6af0b4 new material courtesy of Norbert Voelker
paulson
parents: 15064
diff changeset
  2749
lemma remdups_eq_nil_right_iff [simp]: "([] = remdups x) = (x = [])"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  2750
by (induct x, auto)
15072
4861bf6af0b4 new material courtesy of Norbert Voelker
paulson
parents: 15064
diff changeset
  2751
15245
5a21d9a8f14b Added a few lemmas
nipkow
parents: 15236
diff changeset
  2752
lemma length_remdups_leq[iff]: "length(remdups xs) <= length xs"
5a21d9a8f14b Added a few lemmas
nipkow
parents: 15236
diff changeset
  2753
by (induct xs) auto
5a21d9a8f14b Added a few lemmas
nipkow
parents: 15236
diff changeset
  2754
5a21d9a8f14b Added a few lemmas
nipkow
parents: 15236
diff changeset
  2755
lemma length_remdups_eq[iff]:
5a21d9a8f14b Added a few lemmas
nipkow
parents: 15236
diff changeset
  2756
  "(length (remdups xs) = length xs) = (remdups xs = xs)"
5a21d9a8f14b Added a few lemmas
nipkow
parents: 15236
diff changeset
  2757
apply(induct xs)
5a21d9a8f14b Added a few lemmas
nipkow
parents: 15236
diff changeset
  2758
 apply auto
5a21d9a8f14b Added a few lemmas
nipkow
parents: 15236
diff changeset
  2759
apply(subgoal_tac "length (remdups xs) <= length xs")
5a21d9a8f14b Added a few lemmas
nipkow
parents: 15236
diff changeset
  2760
 apply arith
5a21d9a8f14b Added a few lemmas
nipkow
parents: 15236
diff changeset
  2761
apply(rule length_remdups_leq)
5a21d9a8f14b Added a few lemmas
nipkow
parents: 15236
diff changeset
  2762
done
5a21d9a8f14b Added a few lemmas
nipkow
parents: 15236
diff changeset
  2763
33945
8493ed132fed added remdups_filter lemma
nipkow
parents: 33640
diff changeset
  2764
lemma remdups_filter: "remdups(filter P xs) = filter P (remdups xs)"
8493ed132fed added remdups_filter lemma
nipkow
parents: 33640
diff changeset
  2765
apply(induct xs)
8493ed132fed added remdups_filter lemma
nipkow
parents: 33640
diff changeset
  2766
apply auto
8493ed132fed added remdups_filter lemma
nipkow
parents: 33640
diff changeset
  2767
done
18490
434e34392c40 new lemmas
nipkow
parents: 18451
diff changeset
  2768
434e34392c40 new lemmas
nipkow
parents: 18451
diff changeset
  2769
lemma distinct_map:
434e34392c40 new lemmas
nipkow
parents: 18451
diff changeset
  2770
  "distinct(map f xs) = (distinct xs & inj_on f (set xs))"
434e34392c40 new lemmas
nipkow
parents: 18451
diff changeset
  2771
by (induct xs) auto
434e34392c40 new lemmas
nipkow
parents: 18451
diff changeset
  2772
434e34392c40 new lemmas
nipkow
parents: 18451
diff changeset
  2773
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  2774
lemma distinct_filter [simp]: "distinct xs ==> distinct (filter P xs)"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2775
by (induct xs) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2776
17501
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2777
lemma distinct_upt[simp]: "distinct[i..<j]"
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2778
by (induct j) auto
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2779
32415
1dddf2f64266 got rid of complicated class finite_intvl_succ and defined "upto" directly on int, the only instance of the class.
nipkow
parents: 32078
diff changeset
  2780
lemma distinct_upto[simp]: "distinct[i..j]"
1dddf2f64266 got rid of complicated class finite_intvl_succ and defined "upto" directly on int, the only instance of the class.
nipkow
parents: 32078
diff changeset
  2781
apply(induct i j rule:upto.induct)
1dddf2f64266 got rid of complicated class finite_intvl_succ and defined "upto" directly on int, the only instance of the class.
nipkow
parents: 32078
diff changeset
  2782
apply(subst upto.simps)
1dddf2f64266 got rid of complicated class finite_intvl_succ and defined "upto" directly on int, the only instance of the class.
nipkow
parents: 32078
diff changeset
  2783
apply(simp)
1dddf2f64266 got rid of complicated class finite_intvl_succ and defined "upto" directly on int, the only instance of the class.
nipkow
parents: 32078
diff changeset
  2784
done
1dddf2f64266 got rid of complicated class finite_intvl_succ and defined "upto" directly on int, the only instance of the class.
nipkow
parents: 32078
diff changeset
  2785
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2786
lemma distinct_take[simp]: "distinct xs \<Longrightarrow> distinct (take i xs)"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2787
apply(induct xs arbitrary: i)
17501
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2788
 apply simp
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2789
apply (case_tac i)
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2790
 apply simp_all
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2791
apply(blast dest:in_set_takeD)
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2792
done
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2793
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2794
lemma distinct_drop[simp]: "distinct xs \<Longrightarrow> distinct (drop i xs)"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  2795
apply(induct xs arbitrary: i)
17501
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2796
 apply simp
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2797
apply (case_tac i)
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2798
 apply simp_all
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2799
done
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2800
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2801
lemma distinct_list_update:
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2802
assumes d: "distinct xs" and a: "a \<notin> set xs - {xs!i}"
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2803
shows "distinct (xs[i:=a])"
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2804
proof (cases "i < length xs")
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2805
  case True
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2806
  with a have "a \<notin> set (take i xs @ xs ! i # drop (Suc i) xs) - {xs!i}"
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2807
    apply (drule_tac id_take_nth_drop) by simp
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2808
  with d True show ?thesis
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2809
    apply (simp add: upd_conv_take_nth_drop)
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2810
    apply (drule subst [OF id_take_nth_drop]) apply assumption
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2811
    apply simp apply (cases "a = xs!i") apply simp by blast
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2812
next
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2813
  case False with d show ?thesis by auto
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2814
qed
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2815
31363
7493b571b37d Added theorems about distinct & concat, map & replicate and concat & replicate
hoelzl
parents: 31264
diff changeset
  2816
lemma distinct_concat:
7493b571b37d Added theorems about distinct & concat, map & replicate and concat & replicate
hoelzl
parents: 31264
diff changeset
  2817
  assumes "distinct xs"
7493b571b37d Added theorems about distinct & concat, map & replicate and concat & replicate
hoelzl
parents: 31264
diff changeset
  2818
  and "\<And> ys. ys \<in> set xs \<Longrightarrow> distinct ys"
7493b571b37d Added theorems about distinct & concat, map & replicate and concat & replicate
hoelzl
parents: 31264
diff changeset
  2819
  and "\<And> ys zs. \<lbrakk> ys \<in> set xs ; zs \<in> set xs ; ys \<noteq> zs \<rbrakk> \<Longrightarrow> set ys \<inter> set zs = {}"
7493b571b37d Added theorems about distinct & concat, map & replicate and concat & replicate
hoelzl
parents: 31264
diff changeset
  2820
  shows "distinct (concat xs)"
7493b571b37d Added theorems about distinct & concat, map & replicate and concat & replicate
hoelzl
parents: 31264
diff changeset
  2821
  using assms by (induct xs) auto
17501
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2822
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2823
text {* It is best to avoid this indexed version of distinct, but
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2824
sometimes it is useful. *}
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2825
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  2826
lemma distinct_conv_nth:
17501
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  2827
"distinct xs = (\<forall>i < size xs. \<forall>j < size xs. i \<noteq> j --> xs!i \<noteq> xs!j)"
15251
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15246
diff changeset
  2828
apply (induct xs, simp, simp)
14208
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
  2829
apply (rule iffI, clarsimp)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2830
 apply (case_tac i)
14208
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
  2831
apply (case_tac j, simp)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2832
apply (simp add: set_conv_nth)
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2833
 apply (case_tac j)
24648
1e8053a6d725 metis too slow
paulson
parents: 24645
diff changeset
  2834
apply (clarsimp simp add: set_conv_nth, simp) 
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2835
apply (rule conjI)
24648
1e8053a6d725 metis too slow
paulson
parents: 24645
diff changeset
  2836
(*TOO SLOW
24632
779fc4fcbf8b metis now available in PreList
paulson
parents: 24617
diff changeset
  2837
apply (metis Zero_neq_Suc gr0_conv_Suc in_set_conv_nth lessI less_trans_Suc nth_Cons' nth_Cons_Suc)
24648
1e8053a6d725 metis too slow
paulson
parents: 24645
diff changeset
  2838
*)
1e8053a6d725 metis too slow
paulson
parents: 24645
diff changeset
  2839
 apply (clarsimp simp add: set_conv_nth)
1e8053a6d725 metis too slow
paulson
parents: 24645
diff changeset
  2840
 apply (erule_tac x = 0 in allE, simp)
1e8053a6d725 metis too slow
paulson
parents: 24645
diff changeset
  2841
 apply (erule_tac x = "Suc i" in allE, simp, clarsimp)
25130
d91391a8705b avoid very slow metis invocation;
wenzelm
parents: 25112
diff changeset
  2842
(*TOO SLOW
24632
779fc4fcbf8b metis now available in PreList
paulson
parents: 24617
diff changeset
  2843
apply (metis Suc_Suc_eq lessI less_trans_Suc nth_Cons_Suc)
25130
d91391a8705b avoid very slow metis invocation;
wenzelm
parents: 25112
diff changeset
  2844
*)
d91391a8705b avoid very slow metis invocation;
wenzelm
parents: 25112
diff changeset
  2845
apply (erule_tac x = "Suc i" in allE, simp)
d91391a8705b avoid very slow metis invocation;
wenzelm
parents: 25112
diff changeset
  2846
apply (erule_tac x = "Suc j" in allE, simp)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  2847
done
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  2848
18490
434e34392c40 new lemmas
nipkow
parents: 18451
diff changeset
  2849
lemma nth_eq_iff_index_eq:
434e34392c40 new lemmas
nipkow
parents: 18451
diff changeset
  2850
 "\<lbrakk> distinct xs; i < length xs; j < length xs \<rbrakk> \<Longrightarrow> (xs!i = xs!j) = (i = j)"
434e34392c40 new lemmas
nipkow
parents: 18451
diff changeset
  2851
by(auto simp: distinct_conv_nth)
434e34392c40 new lemmas
nipkow
parents: 18451
diff changeset
  2852
15110
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  2853
lemma distinct_card: "distinct xs ==> card (set xs) = size xs"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  2854
by (induct xs) auto
14388
04f04408b99b lemmas about card (set xs)
kleing
parents: 14343
diff changeset
  2855
15110
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  2856
lemma card_distinct: "card (set xs) = size xs ==> distinct xs"
14388
04f04408b99b lemmas about card (set xs)
kleing
parents: 14343
diff changeset
  2857
proof (induct xs)
04f04408b99b lemmas about card (set xs)
kleing
parents: 14343
diff changeset
  2858
  case Nil thus ?case by simp
04f04408b99b lemmas about card (set xs)
kleing
parents: 14343
diff changeset
  2859
next
04f04408b99b lemmas about card (set xs)
kleing
parents: 14343
diff changeset
  2860
  case (Cons x xs)
04f04408b99b lemmas about card (set xs)
kleing
parents: 14343
diff changeset
  2861
  show ?case
04f04408b99b lemmas about card (set xs)
kleing
parents: 14343
diff changeset
  2862
  proof (cases "x \<in> set xs")
04f04408b99b lemmas about card (set xs)
kleing
parents: 14343
diff changeset
  2863
    case False with Cons show ?thesis by simp
04f04408b99b lemmas about card (set xs)
kleing
parents: 14343
diff changeset
  2864
  next
04f04408b99b lemmas about card (set xs)
kleing
parents: 14343
diff changeset
  2865
    case True with Cons.prems
04f04408b99b lemmas about card (set xs)
kleing
parents: 14343
diff changeset
  2866
    have "card (set xs) = Suc (length xs)" 
04f04408b99b lemmas about card (set xs)
kleing
parents: 14343
diff changeset
  2867
      by (simp add: card_insert_if split: split_if_asm)
04f04408b99b lemmas about card (set xs)
kleing
parents: 14343
diff changeset
  2868
    moreover have "card (set xs) \<le> length xs" by (rule card_length)
04f04408b99b lemmas about card (set xs)
kleing
parents: 14343
diff changeset
  2869
    ultimately have False by simp
04f04408b99b lemmas about card (set xs)
kleing
parents: 14343
diff changeset
  2870
    thus ?thesis ..
04f04408b99b lemmas about card (set xs)
kleing
parents: 14343
diff changeset
  2871
  qed
04f04408b99b lemmas about card (set xs)
kleing
parents: 14343
diff changeset
  2872
qed
04f04408b99b lemmas about card (set xs)
kleing
parents: 14343
diff changeset
  2873
45115
93c1ac6727a3 adding lemma to List library for executable equation of the (refl) transitive closure
bulwahn
parents: 44928
diff changeset
  2874
lemma distinct_length_filter: "distinct xs \<Longrightarrow> length (filter P xs) = card ({x. P x} Int set xs)"
93c1ac6727a3 adding lemma to List library for executable equation of the (refl) transitive closure
bulwahn
parents: 44928
diff changeset
  2875
by (induct xs) (auto)
93c1ac6727a3 adding lemma to List library for executable equation of the (refl) transitive closure
bulwahn
parents: 44928
diff changeset
  2876
25287
094dab519ff5 added lemmas
nipkow
parents: 25277
diff changeset
  2877
lemma not_distinct_decomp: "~ distinct ws ==> EX xs ys zs y. ws = xs@[y]@ys@[y]@zs"
094dab519ff5 added lemmas
nipkow
parents: 25277
diff changeset
  2878
apply (induct n == "length ws" arbitrary:ws) apply simp
094dab519ff5 added lemmas
nipkow
parents: 25277
diff changeset
  2879
apply(case_tac ws) apply simp
094dab519ff5 added lemmas
nipkow
parents: 25277
diff changeset
  2880
apply (simp split:split_if_asm)
094dab519ff5 added lemmas
nipkow
parents: 25277
diff changeset
  2881
apply (metis Cons_eq_appendI eq_Nil_appendI split_list)
094dab519ff5 added lemmas
nipkow
parents: 25277
diff changeset
  2882
done
18490
434e34392c40 new lemmas
nipkow
parents: 18451
diff changeset
  2883
434e34392c40 new lemmas
nipkow
parents: 18451
diff changeset
  2884
lemma length_remdups_concat:
44921
58eef4843641 tuned proofs
huffman
parents: 44916
diff changeset
  2885
  "length (remdups (concat xss)) = card (\<Union>xs\<in>set xss. set xs)"
58eef4843641 tuned proofs
huffman
parents: 44916
diff changeset
  2886
  by (simp add: distinct_card [symmetric])
17906
719364f5179b added 2 lemmas
nipkow
parents: 17877
diff changeset
  2887
33639
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2888
lemma length_remdups_card_conv: "length(remdups xs) = card(set xs)"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2889
proof -
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2890
  have xs: "concat[xs] = xs" by simp
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2891
  from length_remdups_concat[of "[xs]"] show ?thesis unfolding xs by simp
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  2892
qed
17906
719364f5179b added 2 lemmas
nipkow
parents: 17877
diff changeset
  2893
36275
c6ca9e258269 lemmas concerning remdups
haftmann
parents: 36199
diff changeset
  2894
lemma remdups_remdups:
c6ca9e258269 lemmas concerning remdups
haftmann
parents: 36199
diff changeset
  2895
  "remdups (remdups xs) = remdups xs"
c6ca9e258269 lemmas concerning remdups
haftmann
parents: 36199
diff changeset
  2896
  by (induct xs) simp_all
c6ca9e258269 lemmas concerning remdups
haftmann
parents: 36199
diff changeset
  2897
36851
5135adb33157 added lemmas concerning last, butlast, insort
haftmann
parents: 36622
diff changeset
  2898
lemma distinct_butlast:
5135adb33157 added lemmas concerning last, butlast, insort
haftmann
parents: 36622
diff changeset
  2899
  assumes "xs \<noteq> []" and "distinct xs"
5135adb33157 added lemmas concerning last, butlast, insort
haftmann
parents: 36622
diff changeset
  2900
  shows "distinct (butlast xs)"
5135adb33157 added lemmas concerning last, butlast, insort
haftmann
parents: 36622
diff changeset
  2901
proof -
5135adb33157 added lemmas concerning last, butlast, insort
haftmann
parents: 36622
diff changeset
  2902
  from `xs \<noteq> []` obtain ys y where "xs = ys @ [y]" by (cases xs rule: rev_cases) auto
5135adb33157 added lemmas concerning last, butlast, insort
haftmann
parents: 36622
diff changeset
  2903
  with `distinct xs` show ?thesis by simp
5135adb33157 added lemmas concerning last, butlast, insort
haftmann
parents: 36622
diff changeset
  2904
qed
5135adb33157 added lemmas concerning last, butlast, insort
haftmann
parents: 36622
diff changeset
  2905
39728
832c42be723e lemma remdups_map_remdups
haftmann
parents: 39613
diff changeset
  2906
lemma remdups_map_remdups:
832c42be723e lemma remdups_map_remdups
haftmann
parents: 39613
diff changeset
  2907
  "remdups (map f (remdups xs)) = remdups (map f xs)"
832c42be723e lemma remdups_map_remdups
haftmann
parents: 39613
diff changeset
  2908
  by (induct xs) simp_all
832c42be723e lemma remdups_map_remdups
haftmann
parents: 39613
diff changeset
  2909
39915
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  2910
lemma distinct_zipI1:
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  2911
  assumes "distinct xs"
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  2912
  shows "distinct (zip xs ys)"
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  2913
proof (rule zip_obtain_same_length)
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  2914
  fix xs' :: "'a list" and ys' :: "'b list" and n
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  2915
  assume "length xs' = length ys'"
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  2916
  assume "xs' = take n xs"
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  2917
  with assms have "distinct xs'" by simp
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  2918
  with `length xs' = length ys'` show "distinct (zip xs' ys')"
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  2919
    by (induct xs' ys' rule: list_induct2) (auto elim: in_set_zipE)
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  2920
qed
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  2921
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  2922
lemma distinct_zipI2:
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  2923
  assumes "distinct ys"
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  2924
  shows "distinct (zip xs ys)"
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  2925
proof (rule zip_obtain_same_length)
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  2926
  fix xs' :: "'b list" and ys' :: "'a list" and n
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  2927
  assume "length xs' = length ys'"
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  2928
  assume "ys' = take n ys"
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  2929
  with assms have "distinct ys'" by simp
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  2930
  with `length xs' = length ys'` show "distinct (zip xs' ys')"
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  2931
    by (induct xs' ys' rule: list_induct2) (auto elim: in_set_zipE)
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  2932
qed
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  2933
44635
3d046864ebe6 added two lemmas about "distinct" to help Sledgehammer
blanchet
parents: 44619
diff changeset
  2934
(* The next two lemmas help Sledgehammer. *)
3d046864ebe6 added two lemmas about "distinct" to help Sledgehammer
blanchet
parents: 44619
diff changeset
  2935
3d046864ebe6 added two lemmas about "distinct" to help Sledgehammer
blanchet
parents: 44619
diff changeset
  2936
lemma distinct_singleton: "distinct [x]" by simp
3d046864ebe6 added two lemmas about "distinct" to help Sledgehammer
blanchet
parents: 44619
diff changeset
  2937
3d046864ebe6 added two lemmas about "distinct" to help Sledgehammer
blanchet
parents: 44619
diff changeset
  2938
lemma distinct_length_2_or_more:
3d046864ebe6 added two lemmas about "distinct" to help Sledgehammer
blanchet
parents: 44619
diff changeset
  2939
"distinct (a # b # xs) \<longleftrightarrow> (a \<noteq> b \<and> distinct (a # xs) \<and> distinct (b # xs))"
3d046864ebe6 added two lemmas about "distinct" to help Sledgehammer
blanchet
parents: 44619
diff changeset
  2940
by (metis distinct.simps(2) hd.simps hd_in_set list.simps(2) set_ConsD set_rev_mp set_subset_Cons)
3d046864ebe6 added two lemmas about "distinct" to help Sledgehammer
blanchet
parents: 44619
diff changeset
  2941
35115
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
  2942
37605
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  2943
subsubsection {* List summation: @{const listsum} and @{text"\<Sum>"}*}
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  2944
39774
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2945
lemma (in monoid_add) listsum_foldl [code]:
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2946
  "listsum = foldl (op +) 0"
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2947
  by (simp add: listsum_def foldl_foldr1 fun_eq_iff)
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2948
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2949
lemma (in monoid_add) listsum_simps [simp]:
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2950
  "listsum [] = 0"
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2951
  "listsum (x#xs) = x + listsum xs"
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2952
  by (simp_all add: listsum_def)
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2953
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2954
lemma (in monoid_add) listsum_append [simp]:
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2955
  "listsum (xs @ ys) = listsum xs + listsum ys"
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2956
  by (induct xs) (simp_all add: add.assoc)
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2957
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2958
lemma (in comm_monoid_add) listsum_rev [simp]:
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2959
  "listsum (rev xs) = listsum xs"
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2960
  by (simp add: listsum_def [of "rev xs"]) (simp add: listsum_foldl foldr_foldl add.commute)
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2961
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2962
lemma (in comm_monoid_add) listsum_map_remove1:
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2963
  "x \<in> set xs \<Longrightarrow> listsum (map f xs) = f x + listsum (map f (remove1 x xs))"
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2964
  by (induct xs) (auto simp add: ac_simps)
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2965
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2966
lemma (in monoid_add) list_size_conv_listsum:
37605
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  2967
  "list_size f xs = listsum (map f xs) + size xs"
39774
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2968
  by (induct xs) auto
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2969
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2970
lemma (in monoid_add) length_concat:
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2971
  "length (concat xss) = listsum (map length xss)"
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2972
  by (induct xss) simp_all
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2973
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2974
lemma (in monoid_add) listsum_map_filter:
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2975
  assumes "\<And>x. x \<in> set xs \<Longrightarrow> \<not> P x \<Longrightarrow> f x = 0"
37605
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  2976
  shows "listsum (map f (filter P xs)) = listsum (map f xs)"
39774
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2977
  using assms by (induct xs) auto
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2978
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2979
lemma (in monoid_add) distinct_listsum_conv_Setsum:
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2980
  "distinct xs \<Longrightarrow> listsum xs = Setsum (set xs)"
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2981
  by (induct xs) simp_all
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2982
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2983
lemma listsum_eq_0_nat_iff_nat [simp]:
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2984
  "listsum ns = (0::nat) \<longleftrightarrow> (\<forall>n \<in> set ns. n = 0)"
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2985
  by (simp add: listsum_foldl)
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2986
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2987
lemma elem_le_listsum_nat:
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2988
  "k < size ns \<Longrightarrow> ns ! k \<le> listsum (ns::nat list)"
37605
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  2989
apply(induct ns arbitrary: k)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  2990
 apply simp
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44635
diff changeset
  2991
apply(fastforce simp add:nth_Cons split: nat.split)
37605
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  2992
done
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  2993
39774
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2994
lemma listsum_update_nat:
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  2995
  "k<size ns \<Longrightarrow> listsum (ns[k := (n::nat)]) = listsum ns + n - ns ! k"
37605
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  2996
apply(induct ns arbitrary:k)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  2997
 apply (auto split:nat.split)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  2998
apply(drule elem_le_listsum_nat)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  2999
apply arith
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3000
done
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3001
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3002
text{* Some syntactic sugar for summing a function over a list: *}
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3003
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3004
syntax
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3005
  "_listsum" :: "pttrn => 'a list => 'b => 'b"    ("(3SUM _<-_. _)" [0, 51, 10] 10)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3006
syntax (xsymbols)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3007
  "_listsum" :: "pttrn => 'a list => 'b => 'b"    ("(3\<Sum>_\<leftarrow>_. _)" [0, 51, 10] 10)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3008
syntax (HTML output)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3009
  "_listsum" :: "pttrn => 'a list => 'b => 'b"    ("(3\<Sum>_\<leftarrow>_. _)" [0, 51, 10] 10)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3010
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3011
translations -- {* Beware of argument permutation! *}
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3012
  "SUM x<-xs. b" == "CONST listsum (CONST map (%x. b) xs)"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3013
  "\<Sum>x\<leftarrow>xs. b" == "CONST listsum (CONST map (%x. b) xs)"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3014
39774
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3015
lemma (in monoid_add) listsum_triv:
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3016
  "(\<Sum>x\<leftarrow>xs. r) = of_nat (length xs) * r"
37605
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3017
  by (induct xs) (simp_all add: left_distrib)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3018
39774
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3019
lemma (in monoid_add) listsum_0 [simp]:
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3020
  "(\<Sum>x\<leftarrow>xs. 0) = 0"
37605
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3021
  by (induct xs) (simp_all add: left_distrib)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3022
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3023
text{* For non-Abelian groups @{text xs} needs to be reversed on one side: *}
39774
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3024
lemma (in ab_group_add) uminus_listsum_map:
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3025
  "- listsum (map f xs) = listsum (map (uminus \<circ> f) xs)"
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3026
  by (induct xs) simp_all
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3027
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3028
lemma (in comm_monoid_add) listsum_addf:
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3029
  "(\<Sum>x\<leftarrow>xs. f x + g x) = listsum (map f xs) + listsum (map g xs)"
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3030
  by (induct xs) (simp_all add: algebra_simps)
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3031
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3032
lemma (in ab_group_add) listsum_subtractf:
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3033
  "(\<Sum>x\<leftarrow>xs. f x - g x) = listsum (map f xs) - listsum (map g xs)"
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3034
  by (induct xs) (simp_all add: algebra_simps)
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3035
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3036
lemma (in semiring_0) listsum_const_mult:
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3037
  "(\<Sum>x\<leftarrow>xs. c * f x) = c * (\<Sum>x\<leftarrow>xs. f x)"
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3038
  by (induct xs) (simp_all add: algebra_simps)
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3039
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3040
lemma (in semiring_0) listsum_mult_const:
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3041
  "(\<Sum>x\<leftarrow>xs. f x * c) = (\<Sum>x\<leftarrow>xs. f x) * c"
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3042
  by (induct xs) (simp_all add: algebra_simps)
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3043
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3044
lemma (in ordered_ab_group_add_abs) listsum_abs:
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3045
  "\<bar>listsum xs\<bar> \<le> listsum (map abs xs)"
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3046
  by (induct xs) (simp_all add: order_trans [OF abs_triangle_ineq])
37605
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3047
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3048
lemma listsum_mono:
39774
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3049
  fixes f g :: "'a \<Rightarrow> 'b::{monoid_add, ordered_ab_semigroup_add}"
37605
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3050
  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)"
39774
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3051
  by (induct xs) (simp, simp add: add_mono)
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3052
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3053
lemma (in monoid_add) listsum_distinct_conv_setsum_set:
37605
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3054
  "distinct xs \<Longrightarrow> listsum (map f xs) = setsum f (set xs)"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3055
  by (induct xs) simp_all
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3056
39774
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3057
lemma (in monoid_add) interv_listsum_conv_setsum_set_nat:
37605
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3058
  "listsum (map f [m..<n]) = setsum f (set [m..<n])"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3059
  by (simp add: listsum_distinct_conv_setsum_set)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3060
39774
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3061
lemma (in monoid_add) interv_listsum_conv_setsum_set_int:
37605
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3062
  "listsum (map f [k..l]) = setsum f (set [k..l])"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3063
  by (simp add: listsum_distinct_conv_setsum_set)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3064
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3065
text {* General equivalence between @{const listsum} and @{const setsum} *}
39774
30cf9d80939e localized listsum
haftmann
parents: 39728
diff changeset
  3066
lemma (in monoid_add) listsum_setsum_nth:
37605
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3067
  "listsum xs = (\<Sum> i = 0 ..< length xs. xs ! i)"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3068
  using interv_listsum_conv_setsum_set_nat [of "op ! xs" 0 "length xs"] by (simp add: map_nth)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3069
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  3070
34978
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  3071
subsubsection {* @{const insert} *}
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  3072
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  3073
lemma in_set_insert [simp]:
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  3074
  "x \<in> set xs \<Longrightarrow> List.insert x xs = xs"
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  3075
  by (simp add: List.insert_def)
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  3076
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  3077
lemma not_in_set_insert [simp]:
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  3078
  "x \<notin> set xs \<Longrightarrow> List.insert x xs = x # xs"
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  3079
  by (simp add: List.insert_def)
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  3080
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  3081
lemma insert_Nil [simp]:
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  3082
  "List.insert x [] = [x]"
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  3083
  by simp
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  3084
35295
397295fa8387 lemma distinct_insert
haftmann
parents: 35248
diff changeset
  3085
lemma set_insert [simp]:
34978
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  3086
  "set (List.insert x xs) = insert x (set xs)"
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  3087
  by (auto simp add: List.insert_def)
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  3088
35295
397295fa8387 lemma distinct_insert
haftmann
parents: 35248
diff changeset
  3089
lemma distinct_insert [simp]:
397295fa8387 lemma distinct_insert
haftmann
parents: 35248
diff changeset
  3090
  "distinct xs \<Longrightarrow> distinct (List.insert x xs)"
397295fa8387 lemma distinct_insert
haftmann
parents: 35248
diff changeset
  3091
  by (simp add: List.insert_def)
397295fa8387 lemma distinct_insert
haftmann
parents: 35248
diff changeset
  3092
36275
c6ca9e258269 lemmas concerning remdups
haftmann
parents: 36199
diff changeset
  3093
lemma insert_remdups:
c6ca9e258269 lemmas concerning remdups
haftmann
parents: 36199
diff changeset
  3094
  "List.insert x (remdups xs) = remdups (List.insert x xs)"
c6ca9e258269 lemmas concerning remdups
haftmann
parents: 36199
diff changeset
  3095
  by (simp add: List.insert_def)
c6ca9e258269 lemmas concerning remdups
haftmann
parents: 36199
diff changeset
  3096
34978
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  3097
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15307
diff changeset
  3098
subsubsection {* @{text remove1} *}
15110
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  3099
18049
156bba334c12 A few new lemmas
nipkow
parents: 17956
diff changeset
  3100
lemma remove1_append:
156bba334c12 A few new lemmas
nipkow
parents: 17956
diff changeset
  3101
  "remove1 x (xs @ ys) =
156bba334c12 A few new lemmas
nipkow
parents: 17956
diff changeset
  3102
  (if x \<in> set xs then remove1 x xs @ ys else xs @ remove1 x ys)"
156bba334c12 A few new lemmas
nipkow
parents: 17956
diff changeset
  3103
by (induct xs) auto
156bba334c12 A few new lemmas
nipkow
parents: 17956
diff changeset
  3104
36903
489c1fbbb028 Multiset: renamed, added and tuned lemmas;
nipkow
parents: 36851
diff changeset
  3105
lemma remove1_commute: "remove1 x (remove1 y zs) = remove1 y (remove1 x zs)"
489c1fbbb028 Multiset: renamed, added and tuned lemmas;
nipkow
parents: 36851
diff changeset
  3106
by (induct zs) auto
489c1fbbb028 Multiset: renamed, added and tuned lemmas;
nipkow
parents: 36851
diff changeset
  3107
23479
10adbdcdc65b new lemmas
nipkow
parents: 23388
diff changeset
  3108
lemma in_set_remove1[simp]:
10adbdcdc65b new lemmas
nipkow
parents: 23388
diff changeset
  3109
  "a \<noteq> b \<Longrightarrow> a : set(remove1 b xs) = (a : set xs)"
10adbdcdc65b new lemmas
nipkow
parents: 23388
diff changeset
  3110
apply (induct xs)
10adbdcdc65b new lemmas
nipkow
parents: 23388
diff changeset
  3111
apply auto
10adbdcdc65b new lemmas
nipkow
parents: 23388
diff changeset
  3112
done
10adbdcdc65b new lemmas
nipkow
parents: 23388
diff changeset
  3113
15110
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  3114
lemma set_remove1_subset: "set(remove1 x xs) <= set xs"
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  3115
apply(induct xs)
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  3116
 apply simp
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  3117
apply simp
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  3118
apply blast
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  3119
done
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  3120
17724
e969fc0a4925 simprules need names
paulson
parents: 17629
diff changeset
  3121
lemma set_remove1_eq [simp]: "distinct xs ==> set(remove1 x xs) = set xs - {x}"
15110
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  3122
apply(induct xs)
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  3123
 apply simp
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  3124
apply simp
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  3125
apply blast
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  3126
done
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  3127
23479
10adbdcdc65b new lemmas
nipkow
parents: 23388
diff changeset
  3128
lemma length_remove1:
30128
365ee7319b86 revert some Suc 0 lemmas back to their original forms; added some simp rules for (1::nat)
huffman
parents: 30079
diff changeset
  3129
  "length(remove1 x xs) = (if x : set xs then length xs - 1 else length xs)"
23479
10adbdcdc65b new lemmas
nipkow
parents: 23388
diff changeset
  3130
apply (induct xs)
10adbdcdc65b new lemmas
nipkow
parents: 23388
diff changeset
  3131
 apply (auto dest!:length_pos_if_in_set)
10adbdcdc65b new lemmas
nipkow
parents: 23388
diff changeset
  3132
done
10adbdcdc65b new lemmas
nipkow
parents: 23388
diff changeset
  3133
18049
156bba334c12 A few new lemmas
nipkow
parents: 17956
diff changeset
  3134
lemma remove1_filter_not[simp]:
156bba334c12 A few new lemmas
nipkow
parents: 17956
diff changeset
  3135
  "\<not> P x \<Longrightarrow> remove1 x (filter P xs) = filter P xs"
156bba334c12 A few new lemmas
nipkow
parents: 17956
diff changeset
  3136
by(induct xs) auto
156bba334c12 A few new lemmas
nipkow
parents: 17956
diff changeset
  3137
39073
8520a1f89db1 Add filter_remove1
hoelzl
parents: 38857
diff changeset
  3138
lemma filter_remove1:
8520a1f89db1 Add filter_remove1
hoelzl
parents: 38857
diff changeset
  3139
  "filter Q (remove1 x xs) = remove1 x (filter Q xs)"
8520a1f89db1 Add filter_remove1
hoelzl
parents: 38857
diff changeset
  3140
by (induct xs) auto
8520a1f89db1 Add filter_remove1
hoelzl
parents: 38857
diff changeset
  3141
15110
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  3142
lemma notin_set_remove1[simp]: "x ~: set xs ==> x ~: set(remove1 y xs)"
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  3143
apply(insert set_remove1_subset)
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  3144
apply fast
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  3145
done
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  3146
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  3147
lemma distinct_remove1[simp]: "distinct xs ==> distinct(remove1 x xs)"
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  3148
by (induct xs) simp_all
78b5636eabc7 Added a number of new thms and the new function remove1
nipkow
parents: 15072
diff changeset
  3149
36275
c6ca9e258269 lemmas concerning remdups
haftmann
parents: 36199
diff changeset
  3150
lemma remove1_remdups:
c6ca9e258269 lemmas concerning remdups
haftmann
parents: 36199
diff changeset
  3151
  "distinct xs \<Longrightarrow> remove1 x (remdups xs) = remdups (remove1 x xs)"
c6ca9e258269 lemmas concerning remdups
haftmann
parents: 36199
diff changeset
  3152
  by (induct xs) simp_all
c6ca9e258269 lemmas concerning remdups
haftmann
parents: 36199
diff changeset
  3153
37107
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  3154
lemma remove1_idem:
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  3155
  assumes "x \<notin> set xs"
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  3156
  shows "remove1 x xs = xs"
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  3157
  using assms by (induct xs) simp_all
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  3158
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  3159
27693
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3160
subsubsection {* @{text removeAll} *}
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3161
34978
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  3162
lemma removeAll_filter_not_eq:
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  3163
  "removeAll x = filter (\<lambda>y. x \<noteq> y)"
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  3164
proof
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  3165
  fix xs
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  3166
  show "removeAll x xs = filter (\<lambda>y. x \<noteq> y) xs"
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  3167
    by (induct xs) auto
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  3168
qed
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  3169
27693
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3170
lemma removeAll_append[simp]:
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3171
  "removeAll x (xs @ ys) = removeAll x xs @ removeAll x ys"
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3172
by (induct xs) auto
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3173
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3174
lemma set_removeAll[simp]: "set(removeAll x xs) = set xs - {x}"
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3175
by (induct xs) auto
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3176
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3177
lemma removeAll_id[simp]: "x \<notin> set xs \<Longrightarrow> removeAll x xs = xs"
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3178
by (induct xs) auto
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3179
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3180
(* Needs count:: 'a \<Rightarrow> a' list \<Rightarrow> nat
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3181
lemma length_removeAll:
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3182
  "length(removeAll x xs) = length xs - count x xs"
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3183
*)
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3184
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3185
lemma removeAll_filter_not[simp]:
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3186
  "\<not> P x \<Longrightarrow> removeAll x (filter P xs) = filter P xs"
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3187
by(induct xs) auto
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3188
34978
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  3189
lemma distinct_removeAll:
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  3190
  "distinct xs \<Longrightarrow> distinct (removeAll x xs)"
874150ddd50a canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents: 34942
diff changeset
  3191
  by (simp add: removeAll_filter_not_eq)
27693
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3192
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3193
lemma distinct_remove1_removeAll:
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3194
  "distinct xs ==> remove1 x xs = removeAll x xs"
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3195
by (induct xs) simp_all
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3196
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3197
lemma map_removeAll_inj_on: "inj_on f (insert x (set xs)) \<Longrightarrow>
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3198
  map f (removeAll x xs) = removeAll (f x) (map f xs)"
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3199
by (induct xs) (simp_all add:inj_on_def)
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3200
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3201
lemma map_removeAll_inj: "inj f \<Longrightarrow>
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3202
  map f (removeAll x xs) = removeAll (f x) (map f xs)"
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3203
by(metis map_removeAll_inj_on subset_inj_on subset_UNIV)
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3204
73253a4e3ee2 added removeAll
nipkow
parents: 27381
diff changeset
  3205
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15307
diff changeset
  3206
subsubsection {* @{text replicate} *}
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  3207
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  3208
lemma length_replicate [simp]: "length (replicate n x) = n"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  3209
by (induct n) auto
13124
6e1decd8a7a9 new thm distinct_conv_nth
nipkow
parents: 13122
diff changeset
  3210
36622
e393a91f86df Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
hoelzl
parents: 36275
diff changeset
  3211
lemma Ex_list_of_length: "\<exists>xs. length xs = n"
e393a91f86df Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
hoelzl
parents: 36275
diff changeset
  3212
by (rule exI[of _ "replicate n undefined"]) simp
e393a91f86df Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
hoelzl
parents: 36275
diff changeset
  3213
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  3214
lemma map_replicate [simp]: "map f (replicate n x) = replicate n (f x)"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  3215
by (induct n) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  3216
31363
7493b571b37d Added theorems about distinct & concat, map & replicate and concat & replicate
hoelzl
parents: 31264
diff changeset
  3217
lemma map_replicate_const:
7493b571b37d Added theorems about distinct & concat, map & replicate and concat & replicate
hoelzl
parents: 31264
diff changeset
  3218
  "map (\<lambda> x. k) lst = replicate (length lst) k"
7493b571b37d Added theorems about distinct & concat, map & replicate and concat & replicate
hoelzl
parents: 31264
diff changeset
  3219
  by (induct lst) auto
7493b571b37d Added theorems about distinct & concat, map & replicate and concat & replicate
hoelzl
parents: 31264
diff changeset
  3220
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  3221
lemma replicate_app_Cons_same:
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  3222
"(replicate n x) @ (x # xs) = x # replicate n x @ xs"
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  3223
by (induct n) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  3224
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  3225
lemma rev_replicate [simp]: "rev (replicate n x) = replicate n x"
14208
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
  3226
apply (induct n, simp)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  3227
apply (simp add: replicate_app_Cons_same)
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  3228
done
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  3229
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  3230
lemma replicate_add: "replicate (n + m) x = replicate n x @ replicate m x"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  3231
by (induct n) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  3232
16397
c047008f88d4 added lemmas
nipkow
parents: 15870
diff changeset
  3233
text{* Courtesy of Matthias Daum: *}
c047008f88d4 added lemmas
nipkow
parents: 15870
diff changeset
  3234
lemma append_replicate_commute:
c047008f88d4 added lemmas
nipkow
parents: 15870
diff changeset
  3235
  "replicate n x @ replicate k x = replicate k x @ replicate n x"
c047008f88d4 added lemmas
nipkow
parents: 15870
diff changeset
  3236
apply (simp add: replicate_add [THEN sym])
c047008f88d4 added lemmas
nipkow
parents: 15870
diff changeset
  3237
apply (simp add: add_commute)
c047008f88d4 added lemmas
nipkow
parents: 15870
diff changeset
  3238
done
c047008f88d4 added lemmas
nipkow
parents: 15870
diff changeset
  3239
31080
21ffc770ebc0 lemmas by Andreas Lochbihler
nipkow
parents: 31077
diff changeset
  3240
text{* Courtesy of Andreas Lochbihler: *}
21ffc770ebc0 lemmas by Andreas Lochbihler
nipkow
parents: 31077
diff changeset
  3241
lemma filter_replicate:
21ffc770ebc0 lemmas by Andreas Lochbihler
nipkow
parents: 31077
diff changeset
  3242
  "filter P (replicate n x) = (if P x then replicate n x else [])"
21ffc770ebc0 lemmas by Andreas Lochbihler
nipkow
parents: 31077
diff changeset
  3243
by(induct n) auto
21ffc770ebc0 lemmas by Andreas Lochbihler
nipkow
parents: 31077
diff changeset
  3244
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  3245
lemma hd_replicate [simp]: "n \<noteq> 0 ==> hd (replicate n x) = x"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  3246
by (induct n) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  3247
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  3248
lemma tl_replicate [simp]: "n \<noteq> 0 ==> tl (replicate n x) = replicate (n - 1) x"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  3249
by (induct n) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  3250
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  3251
lemma last_replicate [simp]: "n \<noteq> 0 ==> last (replicate n x) = x"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  3252
by (atomize (full), induct n) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  3253
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  3254
lemma nth_replicate[simp]: "i < n ==> (replicate n x)!i = x"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  3255
apply (induct n arbitrary: i, simp)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  3256
apply (simp add: nth_Cons split: nat.split)
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  3257
done
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  3258
16397
c047008f88d4 added lemmas
nipkow
parents: 15870
diff changeset
  3259
text{* Courtesy of Matthias Daum (2 lemmas): *}
c047008f88d4 added lemmas
nipkow
parents: 15870
diff changeset
  3260
lemma take_replicate[simp]: "take i (replicate k x) = replicate (min i k) x"
c047008f88d4 added lemmas
nipkow
parents: 15870
diff changeset
  3261
apply (case_tac "k \<le> i")
c047008f88d4 added lemmas
nipkow
parents: 15870
diff changeset
  3262
 apply  (simp add: min_def)
c047008f88d4 added lemmas
nipkow
parents: 15870
diff changeset
  3263
apply (drule not_leE)
c047008f88d4 added lemmas
nipkow
parents: 15870
diff changeset
  3264
apply (simp add: min_def)
c047008f88d4 added lemmas
nipkow
parents: 15870
diff changeset
  3265
apply (subgoal_tac "replicate k x = replicate i x @ replicate (k - i) x")
c047008f88d4 added lemmas
nipkow
parents: 15870
diff changeset
  3266
 apply  simp
c047008f88d4 added lemmas
nipkow
parents: 15870
diff changeset
  3267
apply (simp add: replicate_add [symmetric])
c047008f88d4 added lemmas
nipkow
parents: 15870
diff changeset
  3268
done
c047008f88d4 added lemmas
nipkow
parents: 15870
diff changeset
  3269
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  3270
lemma drop_replicate[simp]: "drop i (replicate k x) = replicate (k-i) x"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  3271
apply (induct k arbitrary: i)
16397
c047008f88d4 added lemmas
nipkow
parents: 15870
diff changeset
  3272
 apply simp
c047008f88d4 added lemmas
nipkow
parents: 15870
diff changeset
  3273
apply clarsimp
c047008f88d4 added lemmas
nipkow
parents: 15870
diff changeset
  3274
apply (case_tac i)
c047008f88d4 added lemmas
nipkow
parents: 15870
diff changeset
  3275
 apply simp
c047008f88d4 added lemmas
nipkow
parents: 15870
diff changeset
  3276
apply clarsimp
c047008f88d4 added lemmas
nipkow
parents: 15870
diff changeset
  3277
done
c047008f88d4 added lemmas
nipkow
parents: 15870
diff changeset
  3278
c047008f88d4 added lemmas
nipkow
parents: 15870
diff changeset
  3279
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  3280
lemma set_replicate_Suc: "set (replicate (Suc n) x) = {x}"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  3281
by (induct n) auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  3282
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  3283
lemma set_replicate [simp]: "n \<noteq> 0 ==> set (replicate n x) = {x}"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  3284
by (fast dest!: not0_implies_Suc intro!: set_replicate_Suc)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  3285
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  3286
lemma set_replicate_conv_if: "set (replicate n x) = (if n = 0 then {} else {x})"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  3287
by auto
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  3288
37456
0a1cc2675958 tuned set_replicate lemmas
nipkow
parents: 37455
diff changeset
  3289
lemma in_set_replicate[simp]: "(x : set (replicate n y)) = (x = y & n \<noteq> 0)"
0a1cc2675958 tuned set_replicate lemmas
nipkow
parents: 37455
diff changeset
  3290
by (simp add: set_replicate_conv_if)
0a1cc2675958 tuned set_replicate lemmas
nipkow
parents: 37455
diff changeset
  3291
37454
9132a5955127 added lemmas
nipkow
parents: 37424
diff changeset
  3292
lemma Ball_set_replicate[simp]:
9132a5955127 added lemmas
nipkow
parents: 37424
diff changeset
  3293
  "(ALL x : set(replicate n a). P x) = (P a | n=0)"
9132a5955127 added lemmas
nipkow
parents: 37424
diff changeset
  3294
by(simp add: set_replicate_conv_if)
9132a5955127 added lemmas
nipkow
parents: 37424
diff changeset
  3295
9132a5955127 added lemmas
nipkow
parents: 37424
diff changeset
  3296
lemma Bex_set_replicate[simp]:
9132a5955127 added lemmas
nipkow
parents: 37424
diff changeset
  3297
  "(EX x : set(replicate n a). P x) = (P a & n\<noteq>0)"
9132a5955127 added lemmas
nipkow
parents: 37424
diff changeset
  3298
by(simp add: set_replicate_conv_if)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  3299
24796
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  3300
lemma replicate_append_same:
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  3301
  "replicate i x @ [x] = x # replicate i x"
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  3302
  by (induct i) simp_all
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  3303
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  3304
lemma map_replicate_trivial:
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  3305
  "map (\<lambda>i. x) [0..<i] = replicate i x"
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  3306
  by (induct i) (simp_all add: replicate_append_same)
529e458f84d2 added some lemmas
haftmann
parents: 24748
diff changeset
  3307
31363
7493b571b37d Added theorems about distinct & concat, map & replicate and concat & replicate
hoelzl
parents: 31264
diff changeset
  3308
lemma concat_replicate_trivial[simp]:
7493b571b37d Added theorems about distinct & concat, map & replicate and concat & replicate
hoelzl
parents: 31264
diff changeset
  3309
  "concat (replicate i []) = []"
7493b571b37d Added theorems about distinct & concat, map & replicate and concat & replicate
hoelzl
parents: 31264
diff changeset
  3310
  by (induct i) (auto simp add: map_replicate_const)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  3311
28642
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3312
lemma replicate_empty[simp]: "(replicate n x = []) \<longleftrightarrow> n=0"
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3313
by (induct n) auto
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3314
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3315
lemma empty_replicate[simp]: "([] = replicate n x) \<longleftrightarrow> n=0"
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3316
by (induct n) auto
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3317
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3318
lemma replicate_eq_replicate[simp]:
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3319
  "(replicate m x = replicate n y) \<longleftrightarrow> (m=n & (m\<noteq>0 \<longrightarrow> x=y))"
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3320
apply(induct m arbitrary: n)
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3321
 apply simp
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3322
apply(induct_tac n)
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3323
apply auto
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3324
done
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3325
39534
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3326
lemma replicate_length_filter:
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3327
  "replicate (length (filter (\<lambda>y. x = y) xs)) x = filter (\<lambda>y. x = y) xs"
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3328
  by (induct xs) auto
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3329
42714
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3330
lemma comm_append_are_replicate:
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3331
  fixes xs ys :: "'a list"
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3332
  assumes "xs \<noteq> []" "ys \<noteq> []"
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3333
  assumes "xs @ ys = ys @ xs"
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3334
  shows "\<exists>m n zs. concat (replicate m zs) = xs \<and> concat (replicate n zs) = ys"
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3335
  using assms
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3336
proof (induct "length (xs @ ys)" arbitrary: xs ys rule: less_induct)
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3337
  case less
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3338
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3339
  def xs' \<equiv> "if (length xs \<le> length ys) then xs else ys"
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3340
    and ys' \<equiv> "if (length xs \<le> length ys) then ys else xs"
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3341
  then have
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3342
    prems': "length xs' \<le> length ys'"
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3343
            "xs' @ ys' = ys' @ xs'"
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3344
      and "xs' \<noteq> []"
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3345
      and len: "length (xs @ ys) = length (xs' @ ys')"
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3346
    using less by (auto intro: less.hyps)
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3347
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3348
  from prems'
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3349
  obtain ws where "ys' = xs' @ ws"
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3350
    by (auto simp: append_eq_append_conv2)
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3351
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3352
  have "\<exists>m n zs. concat (replicate m zs) = xs' \<and> concat (replicate n zs) = ys'"
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3353
  proof (cases "ws = []")
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3354
    case True
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3355
    then have "concat (replicate 1 xs') = xs'"
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3356
      and "concat (replicate 1 xs') = ys'"
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3357
      using `ys' = xs' @ ws` by auto
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3358
    then show ?thesis by blast
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3359
  next
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3360
    case False
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3361
    from `ys' = xs' @ ws` and `xs' @ ys' = ys' @ xs'`
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3362
    have "xs' @ ws = ws @ xs'" by simp
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3363
    then have "\<exists>m n zs. concat (replicate m zs) = xs' \<and> concat (replicate n zs) = ws"
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3364
      using False and `xs' \<noteq> []` and `ys' = xs' @ ws` and len
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3365
      by (intro less.hyps) auto
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3366
    then obtain m n zs where "concat (replicate m zs) = xs'"
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3367
      and "concat (replicate n zs) = ws" by blast
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3368
    moreover
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3369
    then have "concat (replicate (m + n) zs) = ys'"
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3370
      using `ys' = xs' @ ws`
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3371
      by (simp add: replicate_add)
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3372
    ultimately
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3373
    show ?thesis by blast
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3374
  qed
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3375
  then show ?case
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3376
    using xs'_def ys'_def by metis
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3377
qed
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3378
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3379
lemma comm_append_is_replicate:
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3380
  fixes xs ys :: "'a list"
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3381
  assumes "xs \<noteq> []" "ys \<noteq> []"
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3382
  assumes "xs @ ys = ys @ xs"
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3383
  shows "\<exists>n zs. n > 1 \<and> concat (replicate n zs) = xs @ ys"
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3384
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3385
proof -
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3386
  obtain m n zs where "concat (replicate m zs) = xs"
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3387
    and "concat (replicate n zs) = ys"
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3388
    using assms by (metis comm_append_are_replicate)
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3389
  then have "m + n > 1" and "concat (replicate (m+n) zs) = xs @ ys"
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3390
    using `xs \<noteq> []` and `ys \<noteq> []`
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3391
    by (auto simp: replicate_add)
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3392
  then show ?thesis by blast
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3393
qed
fcba668b0839 add a lemma about commutative append to List.thy
noschinl
parents: 42713
diff changeset
  3394
28642
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3395
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15307
diff changeset
  3396
subsubsection{*@{text rotate1} and @{text rotate}*}
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3397
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3398
lemma rotate_simps[simp]: "rotate1 [] = [] \<and> rotate1 (x#xs) = xs @ [x]"
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3399
by(simp add:rotate1_def)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3400
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3401
lemma rotate0[simp]: "rotate 0 = id"
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3402
by(simp add:rotate_def)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3403
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3404
lemma rotate_Suc[simp]: "rotate (Suc n) xs = rotate1(rotate n xs)"
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3405
by(simp add:rotate_def)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3406
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3407
lemma rotate_add:
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3408
  "rotate (m+n) = rotate m o rotate n"
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3409
by(simp add:rotate_def funpow_add)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3410
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3411
lemma rotate_rotate: "rotate m (rotate n xs) = rotate (m+n) xs"
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3412
by(simp add:rotate_add)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3413
18049
156bba334c12 A few new lemmas
nipkow
parents: 17956
diff changeset
  3414
lemma rotate1_rotate_swap: "rotate1 (rotate n xs) = rotate n (rotate1 xs)"
156bba334c12 A few new lemmas
nipkow
parents: 17956
diff changeset
  3415
by(simp add:rotate_def funpow_swap1)
156bba334c12 A few new lemmas
nipkow
parents: 17956
diff changeset
  3416
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3417
lemma rotate1_length01[simp]: "length xs <= 1 \<Longrightarrow> rotate1 xs = xs"
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3418
by(cases xs) simp_all
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3419
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3420
lemma rotate_length01[simp]: "length xs <= 1 \<Longrightarrow> rotate n xs = xs"
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3421
apply(induct n)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3422
 apply simp
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3423
apply (simp add:rotate_def)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  3424
done
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  3425
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3426
lemma rotate1_hd_tl: "xs \<noteq> [] \<Longrightarrow> rotate1 xs = tl xs @ [hd xs]"
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3427
by(simp add:rotate1_def split:list.split)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3428
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3429
lemma rotate_drop_take:
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3430
  "rotate n xs = drop (n mod length xs) xs @ take (n mod length xs) xs"
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3431
apply(induct n)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3432
 apply simp
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3433
apply(simp add:rotate_def)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3434
apply(cases "xs = []")
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3435
 apply (simp)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3436
apply(case_tac "n mod length xs = 0")
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3437
 apply(simp add:mod_Suc)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3438
 apply(simp add: rotate1_hd_tl drop_Suc take_Suc)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3439
apply(simp add:mod_Suc rotate1_hd_tl drop_Suc[symmetric] drop_tl[symmetric]
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3440
                take_hd_drop linorder_not_le)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  3441
done
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  3442
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3443
lemma rotate_conv_mod: "rotate n xs = rotate (n mod length xs) xs"
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3444
by(simp add:rotate_drop_take)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3445
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3446
lemma rotate_id[simp]: "n mod length xs = 0 \<Longrightarrow> rotate n xs = xs"
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3447
by(simp add:rotate_drop_take)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3448
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3449
lemma length_rotate1[simp]: "length(rotate1 xs) = length xs"
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3450
by(simp add:rotate1_def split:list.split)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3451
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  3452
lemma length_rotate[simp]: "length(rotate n xs) = length xs"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  3453
by (induct n arbitrary: xs) (simp_all add:rotate_def)
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3454
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3455
lemma distinct1_rotate[simp]: "distinct(rotate1 xs) = distinct xs"
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3456
by(simp add:rotate1_def split:list.split) blast
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3457
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3458
lemma distinct_rotate[simp]: "distinct(rotate n xs) = distinct xs"
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3459
by (induct n) (simp_all add:rotate_def)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3460
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3461
lemma rotate_map: "rotate n (map f xs) = map f (rotate n xs)"
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3462
by(simp add:rotate_drop_take take_map drop_map)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3463
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3464
lemma set_rotate1[simp]: "set(rotate1 xs) = set xs"
41463
edbf0a86fb1c adding simproc to rewrite list comprehensions to set comprehensions; adopting proofs
bulwahn
parents: 41372
diff changeset
  3465
by (cases xs) (auto simp add:rotate1_def)
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3466
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3467
lemma set_rotate[simp]: "set(rotate n xs) = set xs"
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3468
by (induct n) (simp_all add:rotate_def)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3469
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3470
lemma rotate1_is_Nil_conv[simp]: "(rotate1 xs = []) = (xs = [])"
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3471
by(simp add:rotate1_def split:list.split)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3472
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3473
lemma rotate_is_Nil_conv[simp]: "(rotate n xs = []) = (xs = [])"
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  3474
by (induct n) (simp_all add:rotate_def)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  3475
15439
71c0f98e31f1 made diff_less a simp rule
nipkow
parents: 15426
diff changeset
  3476
lemma rotate_rev:
71c0f98e31f1 made diff_less a simp rule
nipkow
parents: 15426
diff changeset
  3477
  "rotate n (rev xs) = rev(rotate (length xs - (n mod length xs)) xs)"
71c0f98e31f1 made diff_less a simp rule
nipkow
parents: 15426
diff changeset
  3478
apply(simp add:rotate_drop_take rev_drop rev_take)
71c0f98e31f1 made diff_less a simp rule
nipkow
parents: 15426
diff changeset
  3479
apply(cases "length xs = 0")
71c0f98e31f1 made diff_less a simp rule
nipkow
parents: 15426
diff changeset
  3480
 apply simp
71c0f98e31f1 made diff_less a simp rule
nipkow
parents: 15426
diff changeset
  3481
apply(cases "n mod length xs = 0")
71c0f98e31f1 made diff_less a simp rule
nipkow
parents: 15426
diff changeset
  3482
 apply simp
71c0f98e31f1 made diff_less a simp rule
nipkow
parents: 15426
diff changeset
  3483
apply(simp add:rotate_drop_take rev_drop rev_take)
71c0f98e31f1 made diff_less a simp rule
nipkow
parents: 15426
diff changeset
  3484
done
71c0f98e31f1 made diff_less a simp rule
nipkow
parents: 15426
diff changeset
  3485
18423
d7859164447f new lemmas
nipkow
parents: 18336
diff changeset
  3486
lemma hd_rotate_conv_nth: "xs \<noteq> [] \<Longrightarrow> hd(rotate n xs) = xs!(n mod length xs)"
d7859164447f new lemmas
nipkow
parents: 18336
diff changeset
  3487
apply(simp add:rotate_drop_take hd_append hd_drop_conv_nth hd_conv_nth)
d7859164447f new lemmas
nipkow
parents: 18336
diff changeset
  3488
apply(subgoal_tac "length xs \<noteq> 0")
d7859164447f new lemmas
nipkow
parents: 18336
diff changeset
  3489
 prefer 2 apply simp
d7859164447f new lemmas
nipkow
parents: 18336
diff changeset
  3490
using mod_less_divisor[of "length xs" n] by arith
d7859164447f new lemmas
nipkow
parents: 18336
diff changeset
  3491
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  3492
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15307
diff changeset
  3493
subsubsection {* @{text sublist} --- a generalization of @{text nth} to sets *}
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  3494
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  3495
lemma sublist_empty [simp]: "sublist xs {} = []"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  3496
by (auto simp add: sublist_def)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  3497
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  3498
lemma sublist_nil [simp]: "sublist [] A = []"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  3499
by (auto simp add: sublist_def)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  3500
15281
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  3501
lemma length_sublist:
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  3502
  "length(sublist xs I) = card{i. i < length xs \<and> i : I}"
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  3503
by(simp add: sublist_def length_filter_conv_card cong:conj_cong)
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  3504
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  3505
lemma sublist_shift_lemma_Suc:
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  3506
  "map fst (filter (%p. P(Suc(snd p))) (zip xs is)) =
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  3507
   map fst (filter (%p. P(snd p)) (zip xs (map Suc is)))"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  3508
apply(induct xs arbitrary: "is")
15281
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  3509
 apply simp
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  3510
apply (case_tac "is")
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  3511
 apply simp
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  3512
apply simp
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  3513
done
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  3514
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  3515
lemma sublist_shift_lemma:
23279
e39dd93161d9 tuned list comprehension, changed filter syntax from : to <-
nipkow
parents: 23246
diff changeset
  3516
     "map fst [p<-zip xs [i..<i + length xs] . snd p : A] =
e39dd93161d9 tuned list comprehension, changed filter syntax from : to <-
nipkow
parents: 23246
diff changeset
  3517
      map fst [p<-zip xs [0..<length xs] . snd p + i : A]"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  3518
by (induct xs rule: rev_induct) (simp_all add: add_commute)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  3519
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  3520
lemma sublist_append:
15168
33a08cfc3ae5 new functions for sets of lists
paulson
parents: 15140
diff changeset
  3521
     "sublist (l @ l') A = sublist l A @ sublist l' {j. j + length l : A}"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  3522
apply (unfold sublist_def)
14208
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
  3523
apply (induct l' rule: rev_induct, simp)
44921
58eef4843641 tuned proofs
huffman
parents: 44916
diff changeset
  3524
apply (simp add: upt_add_eq_append[of 0] sublist_shift_lemma)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  3525
apply (simp add: add_commute)
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  3526
done
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  3527
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  3528
lemma sublist_Cons:
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  3529
"sublist (x # l) A = (if 0:A then [x] else []) @ sublist l {j. Suc j : A}"
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  3530
apply (induct l rule: rev_induct)
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  3531
 apply (simp add: sublist_def)
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  3532
apply (simp del: append_Cons add: append_Cons[symmetric] sublist_append)
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  3533
done
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  3534
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  3535
lemma set_sublist: "set(sublist xs I) = {xs!i|i. i<size xs \<and> i \<in> I}"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  3536
apply(induct xs arbitrary: I)
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25157
diff changeset
  3537
apply(auto simp: sublist_Cons nth_Cons split:nat.split dest!: gr0_implies_Suc)
15281
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  3538
done
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  3539
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  3540
lemma set_sublist_subset: "set(sublist xs I) \<subseteq> set xs"
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  3541
by(auto simp add:set_sublist)
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  3542
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  3543
lemma notin_set_sublistI[simp]: "x \<notin> set xs \<Longrightarrow> x \<notin> set(sublist xs I)"
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  3544
by(auto simp add:set_sublist)
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  3545
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  3546
lemma in_set_sublistD: "x \<in> set(sublist xs I) \<Longrightarrow> x \<in> set xs"
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  3547
by(auto simp add:set_sublist)
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  3548
13142
1ebd8ed5a1a0 tuned document;
wenzelm
parents: 13124
diff changeset
  3549
lemma sublist_singleton [simp]: "sublist [x] A = (if 0 : A then [x] else [])"
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  3550
by (simp add: sublist_Cons)
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  3551
15281
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  3552
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  3553
lemma distinct_sublistI[simp]: "distinct xs \<Longrightarrow> distinct(sublist xs I)"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  3554
apply(induct xs arbitrary: I)
15281
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  3555
 apply simp
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  3556
apply(auto simp add:sublist_Cons)
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  3557
done
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  3558
bd4611956c7b More lemmas
nipkow
parents: 15251
diff changeset
  3559
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 14981
diff changeset
  3560
lemma sublist_upt_eq_take [simp]: "sublist l {..<n} = take n l"
14208
144f45277d5a misc tidying
paulson
parents: 14187
diff changeset
  3561
apply (induct l rule: rev_induct, simp)
13145
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  3562
apply (simp split: nat_diff_split add: sublist_append)
59bc43b51aa2 *** empty log message ***
nipkow
parents: 13142
diff changeset
  3563
done
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  3564
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  3565
lemma filter_in_sublist:
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  3566
 "distinct xs \<Longrightarrow> filter (%x. x \<in> set(sublist xs s)) xs = sublist xs s"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  3567
proof (induct xs arbitrary: s)
17501
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  3568
  case Nil thus ?case by simp
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  3569
next
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  3570
  case (Cons a xs)
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  3571
  moreover hence "!x. x: set xs \<longrightarrow> x \<noteq> a" by auto
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  3572
  ultimately show ?case by(simp add: sublist_Cons cong:filter_cong)
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  3573
qed
acbebb72e85a added a number of lemmas
nipkow
parents: 17090
diff changeset
  3574
13114
f2b00262bdfc converted;
wenzelm
parents: 12887
diff changeset
  3575
19390
6c7383f80ad1 Added function "splice"
nipkow
parents: 19363
diff changeset
  3576
subsubsection {* @{const splice} *}
6c7383f80ad1 Added function "splice"
nipkow
parents: 19363
diff changeset
  3577
40593
1e57b18d27b1 code eqn for slice was missing; redefined splice with fun
nipkow
parents: 40365
diff changeset
  3578
lemma splice_Nil2 [simp, code]: "splice xs [] = xs"
19390
6c7383f80ad1 Added function "splice"
nipkow
parents: 19363
diff changeset
  3579
by (cases xs) simp_all
6c7383f80ad1 Added function "splice"
nipkow
parents: 19363
diff changeset
  3580
40593
1e57b18d27b1 code eqn for slice was missing; redefined splice with fun
nipkow
parents: 40365
diff changeset
  3581
declare splice.simps(1,3)[code]
1e57b18d27b1 code eqn for slice was missing; redefined splice with fun
nipkow
parents: 40365
diff changeset
  3582
declare splice.simps(2)[simp del]
19390
6c7383f80ad1 Added function "splice"
nipkow
parents: 19363
diff changeset
  3583
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  3584
lemma length_splice[simp]: "length(splice xs ys) = length xs + length ys"
40593
1e57b18d27b1 code eqn for slice was missing; redefined splice with fun
nipkow
parents: 40365
diff changeset
  3585
by (induct xs ys rule: splice.induct) auto
22793
dc13dfd588b2 new lemma splice_length
nipkow
parents: 22633
diff changeset
  3586
35115
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
  3587
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
  3588
subsubsection {* Transpose *}
34933
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3589
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3590
function transpose where
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3591
"transpose []             = []" |
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3592
"transpose ([]     # xss) = transpose xss" |
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3593
"transpose ((x#xs) # xss) =
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3594
  (x # [h. (h#t) \<leftarrow> xss]) # transpose (xs # [t. (h#t) \<leftarrow> xss])"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3595
by pat_completeness auto
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3596
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3597
lemma transpose_aux_filter_head:
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3598
  "concat (map (list_case [] (\<lambda>h t. [h])) xss) =
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3599
  map (\<lambda>xs. hd xs) [ys\<leftarrow>xss . ys \<noteq> []]"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3600
  by (induct xss) (auto split: list.split)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3601
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3602
lemma transpose_aux_filter_tail:
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3603
  "concat (map (list_case [] (\<lambda>h t. [t])) xss) =
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3604
  map (\<lambda>xs. tl xs) [ys\<leftarrow>xss . ys \<noteq> []]"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3605
  by (induct xss) (auto split: list.split)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3606
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3607
lemma transpose_aux_max:
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3608
  "max (Suc (length xs)) (foldr (\<lambda>xs. max (length xs)) xss 0) =
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3609
  Suc (max (length xs) (foldr (\<lambda>x. max (length x - Suc 0)) [ys\<leftarrow>xss . ys\<noteq>[]] 0))"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3610
  (is "max _ ?foldB = Suc (max _ ?foldA)")
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3611
proof (cases "[ys\<leftarrow>xss . ys\<noteq>[]] = []")
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3612
  case True
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3613
  hence "foldr (\<lambda>xs. max (length xs)) xss 0 = 0"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3614
  proof (induct xss)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3615
    case (Cons x xs)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3616
    moreover hence "x = []" by (cases x) auto
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3617
    ultimately show ?case by auto
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3618
  qed simp
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3619
  thus ?thesis using True by simp
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3620
next
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3621
  case False
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3622
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3623
  have foldA: "?foldA = foldr (\<lambda>x. max (length x)) [ys\<leftarrow>xss . ys \<noteq> []] 0 - 1"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3624
    by (induct xss) auto
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3625
  have foldB: "?foldB = foldr (\<lambda>x. max (length x)) [ys\<leftarrow>xss . ys \<noteq> []] 0"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3626
    by (induct xss) auto
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3627
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3628
  have "0 < ?foldB"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3629
  proof -
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3630
    from False
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3631
    obtain z zs where zs: "[ys\<leftarrow>xss . ys \<noteq> []] = z#zs" by (auto simp: neq_Nil_conv)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3632
    hence "z \<in> set ([ys\<leftarrow>xss . ys \<noteq> []])" by auto
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3633
    hence "z \<noteq> []" by auto
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3634
    thus ?thesis
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3635
      unfolding foldB zs
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3636
      by (auto simp: max_def intro: less_le_trans)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3637
  qed
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3638
  thus ?thesis
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3639
    unfolding foldA foldB max_Suc_Suc[symmetric]
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3640
    by simp
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3641
qed
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3642
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3643
termination transpose
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3644
  by (relation "measure (\<lambda>xs. foldr (\<lambda>xs. max (length xs)) xs 0 + length xs)")
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3645
     (auto simp: transpose_aux_filter_tail foldr_map comp_def transpose_aux_max less_Suc_eq_le)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3646
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3647
lemma transpose_empty: "(transpose xs = []) \<longleftrightarrow> (\<forall>x \<in> set xs. x = [])"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3648
  by (induct rule: transpose.induct) simp_all
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3649
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3650
lemma length_transpose:
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3651
  fixes xs :: "'a list list"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3652
  shows "length (transpose xs) = foldr (\<lambda>xs. max (length xs)) xs 0"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3653
  by (induct rule: transpose.induct)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3654
    (auto simp: transpose_aux_filter_tail foldr_map comp_def transpose_aux_max
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3655
                max_Suc_Suc[symmetric] simp del: max_Suc_Suc)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3656
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3657
lemma nth_transpose:
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3658
  fixes xs :: "'a list list"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3659
  assumes "i < length (transpose xs)"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3660
  shows "transpose xs ! i = map (\<lambda>xs. xs ! i) [ys \<leftarrow> xs. i < length ys]"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3661
using assms proof (induct arbitrary: i rule: transpose.induct)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3662
  case (3 x xs xss)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3663
  def XS == "(x # xs) # xss"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3664
  hence [simp]: "XS \<noteq> []" by auto
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3665
  thus ?case
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3666
  proof (cases i)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3667
    case 0
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3668
    thus ?thesis by (simp add: transpose_aux_filter_head hd_conv_nth)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3669
  next
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3670
    case (Suc j)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3671
    have *: "\<And>xss. xs # map tl xss = map tl ((x#xs)#xss)" by simp
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3672
    have **: "\<And>xss. (x#xs) # filter (\<lambda>ys. ys \<noteq> []) xss = filter (\<lambda>ys. ys \<noteq> []) ((x#xs)#xss)" by simp
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3673
    { fix x have "Suc j < length x \<longleftrightarrow> x \<noteq> [] \<and> j < length x - Suc 0"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3674
      by (cases x) simp_all
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3675
    } note *** = this
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3676
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3677
    have j_less: "j < length (transpose (xs # concat (map (list_case [] (\<lambda>h t. [t])) xss)))"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3678
      using "3.prems" by (simp add: transpose_aux_filter_tail length_transpose Suc)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3679
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3680
    show ?thesis
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3681
      unfolding transpose.simps `i = Suc j` nth_Cons_Suc "3.hyps"[OF j_less]
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3682
      apply (auto simp: transpose_aux_filter_tail filter_map comp_def length_transpose * ** *** XS_def[symmetric])
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3683
      apply (rule_tac y=x in list.exhaust)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3684
      by auto
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3685
  qed
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3686
qed simp_all
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3687
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3688
lemma transpose_map_map:
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3689
  "transpose (map (map f) xs) = map (map f) (transpose xs)"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3690
proof (rule nth_equalityI, safe)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3691
  have [simp]: "length (transpose (map (map f) xs)) = length (transpose xs)"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3692
    by (simp add: length_transpose foldr_map comp_def)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3693
  show "length (transpose (map (map f) xs)) = length (map (map f) (transpose xs))" by simp
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3694
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3695
  fix i assume "i < length (transpose (map (map f) xs))"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3696
  thus "transpose (map (map f) xs) ! i = map (map f) (transpose xs) ! i"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3697
    by (simp add: nth_transpose filter_map comp_def)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3698
qed
24616
fac3dd4ade83 sorting
nipkow
parents: 24566
diff changeset
  3699
35115
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
  3700
31557
4e36f2f17c63 two finiteness lemmas by Robert Himmelmann
nipkow
parents: 31455
diff changeset
  3701
subsubsection {* (In)finiteness *}
28642
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3702
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3703
lemma finite_maxlen:
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3704
  "finite (M::'a list set) ==> EX n. ALL s:M. size s < n"
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3705
proof (induct rule: finite.induct)
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3706
  case emptyI show ?case by simp
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3707
next
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3708
  case (insertI M xs)
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3709
  then obtain n where "\<forall>s\<in>M. length s < n" by blast
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3710
  hence "ALL s:insert xs M. size s < max n (size xs) + 1" by auto
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3711
  thus ?case ..
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3712
qed
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3713
45714
ad4242285560 cardinality of sets of lists
hoelzl
parents: 45607
diff changeset
  3714
lemma lists_length_Suc_eq:
ad4242285560 cardinality of sets of lists
hoelzl
parents: 45607
diff changeset
  3715
  "{xs. set xs \<subseteq> A \<and> length xs = Suc n} =
ad4242285560 cardinality of sets of lists
hoelzl
parents: 45607
diff changeset
  3716
    (\<lambda>(xs, n). n#xs) ` ({xs. set xs \<subseteq> A \<and> length xs = n} \<times> A)"
ad4242285560 cardinality of sets of lists
hoelzl
parents: 45607
diff changeset
  3717
  by (auto simp: length_Suc_conv)
ad4242285560 cardinality of sets of lists
hoelzl
parents: 45607
diff changeset
  3718
ad4242285560 cardinality of sets of lists
hoelzl
parents: 45607
diff changeset
  3719
lemma
ad4242285560 cardinality of sets of lists
hoelzl
parents: 45607
diff changeset
  3720
  assumes "finite A"
ad4242285560 cardinality of sets of lists
hoelzl
parents: 45607
diff changeset
  3721
  shows finite_lists_length_eq: "finite {xs. set xs \<subseteq> A \<and> length xs = n}"
ad4242285560 cardinality of sets of lists
hoelzl
parents: 45607
diff changeset
  3722
  and card_lists_length_eq: "card {xs. set xs \<subseteq> A \<and> length xs = n} = (card A)^n"
ad4242285560 cardinality of sets of lists
hoelzl
parents: 45607
diff changeset
  3723
  using `finite A`
ad4242285560 cardinality of sets of lists
hoelzl
parents: 45607
diff changeset
  3724
  by (induct n)
ad4242285560 cardinality of sets of lists
hoelzl
parents: 45607
diff changeset
  3725
     (auto simp: card_image inj_split_Cons lists_length_Suc_eq cong: conj_cong)
31557
4e36f2f17c63 two finiteness lemmas by Robert Himmelmann
nipkow
parents: 31455
diff changeset
  3726
4e36f2f17c63 two finiteness lemmas by Robert Himmelmann
nipkow
parents: 31455
diff changeset
  3727
lemma finite_lists_length_le:
4e36f2f17c63 two finiteness lemmas by Robert Himmelmann
nipkow
parents: 31455
diff changeset
  3728
  assumes "finite A" shows "finite {xs. set xs \<subseteq> A \<and> length xs \<le> n}"
4e36f2f17c63 two finiteness lemmas by Robert Himmelmann
nipkow
parents: 31455
diff changeset
  3729
 (is "finite ?S")
4e36f2f17c63 two finiteness lemmas by Robert Himmelmann
nipkow
parents: 31455
diff changeset
  3730
proof-
4e36f2f17c63 two finiteness lemmas by Robert Himmelmann
nipkow
parents: 31455
diff changeset
  3731
  have "?S = (\<Union>n\<in>{0..n}. {xs. set xs \<subseteq> A \<and> length xs = n})" by auto
4e36f2f17c63 two finiteness lemmas by Robert Himmelmann
nipkow
parents: 31455
diff changeset
  3732
  thus ?thesis by (auto intro: finite_lists_length_eq[OF `finite A`])
4e36f2f17c63 two finiteness lemmas by Robert Himmelmann
nipkow
parents: 31455
diff changeset
  3733
qed
4e36f2f17c63 two finiteness lemmas by Robert Himmelmann
nipkow
parents: 31455
diff changeset
  3734
45714
ad4242285560 cardinality of sets of lists
hoelzl
parents: 45607
diff changeset
  3735
lemma card_lists_length_le:
ad4242285560 cardinality of sets of lists
hoelzl
parents: 45607
diff changeset
  3736
  assumes "finite A" shows "card {xs. set xs \<subseteq> A \<and> length xs \<le> n} = (\<Sum>i\<le>n. card A^i)"
ad4242285560 cardinality of sets of lists
hoelzl
parents: 45607
diff changeset
  3737
proof -
ad4242285560 cardinality of sets of lists
hoelzl
parents: 45607
diff changeset
  3738
  have "(\<Sum>i\<le>n. card A^i) = card (\<Union>i\<le>n. {xs. set xs \<subseteq> A \<and> length xs = i})"
ad4242285560 cardinality of sets of lists
hoelzl
parents: 45607
diff changeset
  3739
    using `finite A`
ad4242285560 cardinality of sets of lists
hoelzl
parents: 45607
diff changeset
  3740
    by (subst card_UN_disjoint)
ad4242285560 cardinality of sets of lists
hoelzl
parents: 45607
diff changeset
  3741
       (auto simp add: card_lists_length_eq finite_lists_length_eq)
ad4242285560 cardinality of sets of lists
hoelzl
parents: 45607
diff changeset
  3742
  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}"
ad4242285560 cardinality of sets of lists
hoelzl
parents: 45607
diff changeset
  3743
    by auto
ad4242285560 cardinality of sets of lists
hoelzl
parents: 45607
diff changeset
  3744
  finally show ?thesis by simp
ad4242285560 cardinality of sets of lists
hoelzl
parents: 45607
diff changeset
  3745
qed
ad4242285560 cardinality of sets of lists
hoelzl
parents: 45607
diff changeset
  3746
28642
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3747
lemma infinite_UNIV_listI: "~ finite(UNIV::'a list set)"
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3748
apply(rule notI)
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3749
apply(drule finite_maxlen)
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3750
apply (metis UNIV_I length_replicate less_not_refl)
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3751
done
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3752
658b598d8af4 added lemmas
nipkow
parents: 28562
diff changeset
  3753
35115
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
  3754
subsection {* Sorting *}
24616
fac3dd4ade83 sorting
nipkow
parents: 24566
diff changeset
  3755
24617
bc484b2671fd sorting
nipkow
parents: 24616
diff changeset
  3756
text{* Currently it is not shown that @{const sort} returns a
bc484b2671fd sorting
nipkow
parents: 24616
diff changeset
  3757
permutation of its input because the nicest proof is via multisets,
bc484b2671fd sorting
nipkow
parents: 24616
diff changeset
  3758
which are not yet available. Alternatively one could define a function
bc484b2671fd sorting
nipkow
parents: 24616
diff changeset
  3759
that counts the number of occurrences of an element in a list and use
bc484b2671fd sorting
nipkow
parents: 24616
diff changeset
  3760
that instead of multisets to state the correctness property. *}
bc484b2671fd sorting
nipkow
parents: 24616
diff changeset
  3761
24616
fac3dd4ade83 sorting
nipkow
parents: 24566
diff changeset
  3762
context linorder
fac3dd4ade83 sorting
nipkow
parents: 24566
diff changeset
  3763
begin
fac3dd4ade83 sorting
nipkow
parents: 24566
diff changeset
  3764
40210
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  3765
lemma length_insort [simp]:
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  3766
  "length (insort_key f x xs) = Suc (length xs)"
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  3767
  by (induct xs) simp_all
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  3768
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  3769
lemma insort_key_left_comm:
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  3770
  assumes "f x \<noteq> f y"
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  3771
  shows "insort_key f y (insort_key f x xs) = insort_key f x (insort_key f y xs)"
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  3772
  by (induct xs) (auto simp add: assms dest: antisym)
33639
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3773
35195
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  3774
lemma insort_left_comm:
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  3775
  "insort x (insort y xs) = insort y (insort x xs)"
40210
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  3776
  by (cases "x = y") (auto intro: insort_key_left_comm)
35195
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  3777
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42809
diff changeset
  3778
lemma comp_fun_commute_insort:
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42809
diff changeset
  3779
  "comp_fun_commute insort"
35195
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  3780
proof
42809
5b45125b15ba use pointfree characterisation for fold_set locale
haftmann
parents: 42714
diff changeset
  3781
qed (simp add: insort_left_comm fun_eq_iff)
35195
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  3782
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  3783
lemma sort_key_simps [simp]:
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  3784
  "sort_key f [] = []"
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  3785
  "sort_key f (x#xs) = insort_key f x (sort_key f xs)"
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  3786
  by (simp_all add: sort_key_def)
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  3787
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  3788
lemma sort_foldl_insort:
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  3789
  "sort xs = foldl (\<lambda>ys x. insort x ys) [] xs"
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  3790
  by (simp add: sort_key_def foldr_foldl foldl_rev insort_left_comm)
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  3791
33639
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3792
lemma length_sort[simp]: "length (sort_key f xs) = length xs"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3793
by (induct xs, auto)
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3794
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24902
diff changeset
  3795
lemma sorted_Cons: "sorted (x#xs) = (sorted xs & (ALL y:set xs. x <= y))"
24616
fac3dd4ade83 sorting
nipkow
parents: 24566
diff changeset
  3796
apply(induct xs arbitrary: x) apply simp
fac3dd4ade83 sorting
nipkow
parents: 24566
diff changeset
  3797
by simp (blast intro: order_trans)
fac3dd4ade83 sorting
nipkow
parents: 24566
diff changeset
  3798
40210
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  3799
lemma sorted_tl:
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  3800
  "sorted xs \<Longrightarrow> sorted (tl xs)"
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  3801
  by (cases xs) (simp_all add: sorted_Cons)
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  3802
24616
fac3dd4ade83 sorting
nipkow
parents: 24566
diff changeset
  3803
lemma sorted_append:
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24902
diff changeset
  3804
  "sorted (xs@ys) = (sorted xs & sorted ys & (\<forall>x \<in> set xs. \<forall>y \<in> set ys. x\<le>y))"
24616
fac3dd4ade83 sorting
nipkow
parents: 24566
diff changeset
  3805
by (induct xs) (auto simp add:sorted_Cons)
fac3dd4ade83 sorting
nipkow
parents: 24566
diff changeset
  3806
31201
3dde56615750 new lemma
nipkow
parents: 31159
diff changeset
  3807
lemma sorted_nth_mono:
33639
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3808
  "sorted xs \<Longrightarrow> i \<le> j \<Longrightarrow> j < length xs \<Longrightarrow> xs!i \<le> xs!j"
31201
3dde56615750 new lemma
nipkow
parents: 31159
diff changeset
  3809
by (induct xs arbitrary: i j) (auto simp:nth_Cons' sorted_Cons)
3dde56615750 new lemma
nipkow
parents: 31159
diff changeset
  3810
33639
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3811
lemma sorted_rev_nth_mono:
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3812
  "sorted (rev xs) \<Longrightarrow> i \<le> j \<Longrightarrow> j < length xs \<Longrightarrow> xs!j \<le> xs!i"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3813
using sorted_nth_mono[ of "rev xs" "length xs - j - 1" "length xs - i - 1"]
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3814
      rev_nth[of "length xs - i - 1" "xs"] rev_nth[of "length xs - j - 1" "xs"]
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3815
by auto
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3816
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3817
lemma sorted_nth_monoI:
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3818
  "(\<And> i j. \<lbrakk> i \<le> j ; j < length xs \<rbrakk> \<Longrightarrow> xs ! i \<le> xs ! j) \<Longrightarrow> sorted xs"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3819
proof (induct xs)
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3820
  case (Cons x xs)
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3821
  have "sorted xs"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3822
  proof (rule Cons.hyps)
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3823
    fix i j assume "i \<le> j" and "j < length xs"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3824
    with Cons.prems[of "Suc i" "Suc j"]
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3825
    show "xs ! i \<le> xs ! j" by auto
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3826
  qed
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3827
  moreover
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3828
  {
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3829
    fix y assume "y \<in> set xs"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3830
    then obtain j where "j < length xs" and "xs ! j = y"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3831
      unfolding in_set_conv_nth by blast
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3832
    with Cons.prems[of 0 "Suc j"]
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3833
    have "x \<le> y"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3834
      by auto
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3835
  }
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3836
  ultimately
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3837
  show ?case
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3838
    unfolding sorted_Cons by auto
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3839
qed simp
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3840
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3841
lemma sorted_equals_nth_mono:
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3842
  "sorted xs = (\<forall>j < length xs. \<forall>i \<le> j. xs ! i \<le> xs ! j)"
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3843
by (auto intro: sorted_nth_monoI sorted_nth_mono)
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3844
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3845
lemma set_insort: "set(insort_key f x xs) = insert x (set xs)"
24616
fac3dd4ade83 sorting
nipkow
parents: 24566
diff changeset
  3846
by (induct xs) auto
fac3dd4ade83 sorting
nipkow
parents: 24566
diff changeset
  3847
33639
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3848
lemma set_sort[simp]: "set(sort_key f xs) = set xs"
24616
fac3dd4ade83 sorting
nipkow
parents: 24566
diff changeset
  3849
by (induct xs) (simp_all add:set_insort)
fac3dd4ade83 sorting
nipkow
parents: 24566
diff changeset
  3850
33639
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3851
lemma distinct_insort: "distinct (insort_key f x xs) = (x \<notin> set xs \<and> distinct xs)"
24616
fac3dd4ade83 sorting
nipkow
parents: 24566
diff changeset
  3852
by(induct xs)(auto simp:set_insort)
fac3dd4ade83 sorting
nipkow
parents: 24566
diff changeset
  3853
33639
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3854
lemma distinct_sort[simp]: "distinct (sort_key f xs) = distinct xs"
44921
58eef4843641 tuned proofs
huffman
parents: 44916
diff changeset
  3855
  by (induct xs) (simp_all add: distinct_insort)
24616
fac3dd4ade83 sorting
nipkow
parents: 24566
diff changeset
  3856
33639
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3857
lemma sorted_insort_key: "sorted (map f (insort_key f x xs)) = sorted (map f xs)"
40210
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  3858
  by (induct xs) (auto simp:sorted_Cons set_insort)
33639
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3859
24616
fac3dd4ade83 sorting
nipkow
parents: 24566
diff changeset
  3860
lemma sorted_insort: "sorted (insort x xs) = sorted xs"
40210
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  3861
  using sorted_insort_key [where f="\<lambda>x. x"] by simp
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  3862
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  3863
theorem sorted_sort_key [simp]: "sorted (map f (sort_key f xs))"
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  3864
  by (induct xs) (auto simp:sorted_insort_key)
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  3865
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  3866
theorem sorted_sort [simp]: "sorted (sort xs)"
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  3867
  using sorted_sort_key [where f="\<lambda>x. x"] by simp
33639
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3868
36851
5135adb33157 added lemmas concerning last, butlast, insort
haftmann
parents: 36622
diff changeset
  3869
lemma sorted_butlast:
5135adb33157 added lemmas concerning last, butlast, insort
haftmann
parents: 36622
diff changeset
  3870
  assumes "xs \<noteq> []" and "sorted xs"
5135adb33157 added lemmas concerning last, butlast, insort
haftmann
parents: 36622
diff changeset
  3871
  shows "sorted (butlast xs)"
5135adb33157 added lemmas concerning last, butlast, insort
haftmann
parents: 36622
diff changeset
  3872
proof -
5135adb33157 added lemmas concerning last, butlast, insort
haftmann
parents: 36622
diff changeset
  3873
  from `xs \<noteq> []` obtain ys y where "xs = ys @ [y]" by (cases xs rule: rev_cases) auto
5135adb33157 added lemmas concerning last, butlast, insort
haftmann
parents: 36622
diff changeset
  3874
  with `sorted xs` show ?thesis by (simp add: sorted_append)
5135adb33157 added lemmas concerning last, butlast, insort
haftmann
parents: 36622
diff changeset
  3875
qed
5135adb33157 added lemmas concerning last, butlast, insort
haftmann
parents: 36622
diff changeset
  3876
  
5135adb33157 added lemmas concerning last, butlast, insort
haftmann
parents: 36622
diff changeset
  3877
lemma insort_not_Nil [simp]:
5135adb33157 added lemmas concerning last, butlast, insort
haftmann
parents: 36622
diff changeset
  3878
  "insort_key f a xs \<noteq> []"
5135adb33157 added lemmas concerning last, butlast, insort
haftmann
parents: 36622
diff changeset
  3879
  by (induct xs) simp_all
5135adb33157 added lemmas concerning last, butlast, insort
haftmann
parents: 36622
diff changeset
  3880
33639
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3881
lemma insort_is_Cons: "\<forall>x\<in>set xs. f a \<le> f x \<Longrightarrow> insort_key f a xs = a # xs"
26143
314c0bcb7df7 Added useful general lemmas from the work with the HeapMonad
bulwahn
parents: 26073
diff changeset
  3882
by (cases xs) auto
314c0bcb7df7 Added useful general lemmas from the work with the HeapMonad
bulwahn
parents: 26073
diff changeset
  3883
44916
840d8c3d9113 added lemma motivated by a more specific lemma in the AFP-KBPs theories
bulwahn
parents: 44890
diff changeset
  3884
lemma sorted_sort_id: "sorted xs \<Longrightarrow> sort xs = xs"
840d8c3d9113 added lemma motivated by a more specific lemma in the AFP-KBPs theories
bulwahn
parents: 44890
diff changeset
  3885
  by (induct xs) (auto simp add: sorted_Cons insort_is_Cons)
840d8c3d9113 added lemma motivated by a more specific lemma in the AFP-KBPs theories
bulwahn
parents: 44890
diff changeset
  3886
39534
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3887
lemma sorted_map_remove1:
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3888
  "sorted (map f xs) \<Longrightarrow> sorted (map f (remove1 x xs))"
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3889
  by (induct xs) (auto simp add: sorted_Cons)
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3890
26143
314c0bcb7df7 Added useful general lemmas from the work with the HeapMonad
bulwahn
parents: 26073
diff changeset
  3891
lemma sorted_remove1: "sorted xs \<Longrightarrow> sorted (remove1 a xs)"
39534
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3892
  using sorted_map_remove1 [of "\<lambda>x. x"] by simp
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3893
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3894
lemma insort_key_remove1:
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3895
  assumes "a \<in> set xs" and "sorted (map f xs)" and "hd (filter (\<lambda>x. f a = f x) xs) = a"
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3896
  shows "insort_key f a (remove1 a xs) = xs"
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3897
using assms proof (induct xs)
33639
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3898
  case (Cons x xs)
39534
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3899
  then show ?case
33639
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3900
  proof (cases "x = a")
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3901
    case False
39534
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3902
    then have "f x \<noteq> f a" using Cons.prems by auto
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3903
    then have "f x < f a" using Cons.prems by (auto simp: sorted_Cons)
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3904
    with `f x \<noteq> f a` show ?thesis using Cons by (auto simp: sorted_Cons insort_is_Cons)
33639
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3905
  qed (auto simp: sorted_Cons insort_is_Cons)
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3906
qed simp
26143
314c0bcb7df7 Added useful general lemmas from the work with the HeapMonad
bulwahn
parents: 26073
diff changeset
  3907
39534
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3908
lemma insort_remove1:
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3909
  assumes "a \<in> set xs" and "sorted xs"
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3910
  shows "insort a (remove1 a xs) = xs"
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3911
proof (rule insort_key_remove1)
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3912
  from `a \<in> set xs` show "a \<in> set xs" .
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3913
  from `sorted xs` show "sorted (map (\<lambda>x. x) xs)" by simp
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3914
  from `a \<in> set xs` have "a \<in> set (filter (op = a) xs)" by auto
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3915
  then have "set (filter (op = a) xs) \<noteq> {}" by auto
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3916
  then have "filter (op = a) xs \<noteq> []" by (auto simp only: set_empty)
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3917
  then have "length (filter (op = a) xs) > 0" by simp
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3918
  then obtain n where n: "Suc n = length (filter (op = a) xs)"
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3919
    by (cases "length (filter (op = a) xs)") simp_all
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3920
  moreover have "replicate (Suc n) a = a # replicate n a"
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3921
    by simp
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3922
  ultimately show "hd (filter (op = a) xs) = a" by (simp add: replicate_length_filter)
c798d4f1b682 generalized lemma insort_remove1 to insort_key_remove1
haftmann
parents: 39302
diff changeset
  3923
qed
26143
314c0bcb7df7 Added useful general lemmas from the work with the HeapMonad
bulwahn
parents: 26073
diff changeset
  3924
314c0bcb7df7 Added useful general lemmas from the work with the HeapMonad
bulwahn
parents: 26073
diff changeset
  3925
lemma sorted_remdups[simp]:
314c0bcb7df7 Added useful general lemmas from the work with the HeapMonad
bulwahn
parents: 26073
diff changeset
  3926
  "sorted l \<Longrightarrow> sorted (remdups l)"
314c0bcb7df7 Added useful general lemmas from the work with the HeapMonad
bulwahn
parents: 26073
diff changeset
  3927
by (induct l) (auto simp: sorted_Cons)
314c0bcb7df7 Added useful general lemmas from the work with the HeapMonad
bulwahn
parents: 26073
diff changeset
  3928
24645
1af302128da2 Generalized [_.._] from nat to linear orders
nipkow
parents: 24640
diff changeset
  3929
lemma sorted_distinct_set_unique:
1af302128da2 Generalized [_.._] from nat to linear orders
nipkow
parents: 24640
diff changeset
  3930
assumes "sorted xs" "distinct xs" "sorted ys" "distinct ys" "set xs = set ys"
1af302128da2 Generalized [_.._] from nat to linear orders
nipkow
parents: 24640
diff changeset
  3931
shows "xs = ys"
1af302128da2 Generalized [_.._] from nat to linear orders
nipkow
parents: 24640
diff changeset
  3932
proof -
26734
a92057c1ee21 dropped some metis calls
haftmann
parents: 26584
diff changeset
  3933
  from assms have 1: "length xs = length ys" by (auto dest!: distinct_card)
24645
1af302128da2 Generalized [_.._] from nat to linear orders
nipkow
parents: 24640
diff changeset
  3934
  from assms show ?thesis
1af302128da2 Generalized [_.._] from nat to linear orders
nipkow
parents: 24640
diff changeset
  3935
  proof(induct rule:list_induct2[OF 1])
1af302128da2 Generalized [_.._] from nat to linear orders
nipkow
parents: 24640
diff changeset
  3936
    case 1 show ?case by simp
1af302128da2 Generalized [_.._] from nat to linear orders
nipkow
parents: 24640
diff changeset
  3937
  next
1af302128da2 Generalized [_.._] from nat to linear orders
nipkow
parents: 24640
diff changeset
  3938
    case 2 thus ?case by (simp add:sorted_Cons)
1af302128da2 Generalized [_.._] from nat to linear orders
nipkow
parents: 24640
diff changeset
  3939
       (metis Diff_insert_absorb antisym insertE insert_iff)
1af302128da2 Generalized [_.._] from nat to linear orders
nipkow
parents: 24640
diff changeset
  3940
  qed
1af302128da2 Generalized [_.._] from nat to linear orders
nipkow
parents: 24640
diff changeset
  3941
qed
1af302128da2 Generalized [_.._] from nat to linear orders
nipkow
parents: 24640
diff changeset
  3942
35603
c0db094d0d80 moved lemma map_sorted_distinct_set_unique
haftmann
parents: 35510
diff changeset
  3943
lemma map_sorted_distinct_set_unique:
c0db094d0d80 moved lemma map_sorted_distinct_set_unique
haftmann
parents: 35510
diff changeset
  3944
  assumes "inj_on f (set xs \<union> set ys)"
c0db094d0d80 moved lemma map_sorted_distinct_set_unique
haftmann
parents: 35510
diff changeset
  3945
  assumes "sorted (map f xs)" "distinct (map f xs)"
c0db094d0d80 moved lemma map_sorted_distinct_set_unique
haftmann
parents: 35510
diff changeset
  3946
    "sorted (map f ys)" "distinct (map f ys)"
c0db094d0d80 moved lemma map_sorted_distinct_set_unique
haftmann
parents: 35510
diff changeset
  3947
  assumes "set xs = set ys"
c0db094d0d80 moved lemma map_sorted_distinct_set_unique
haftmann
parents: 35510
diff changeset
  3948
  shows "xs = ys"
c0db094d0d80 moved lemma map_sorted_distinct_set_unique
haftmann
parents: 35510
diff changeset
  3949
proof -
c0db094d0d80 moved lemma map_sorted_distinct_set_unique
haftmann
parents: 35510
diff changeset
  3950
  from assms have "map f xs = map f ys"
c0db094d0d80 moved lemma map_sorted_distinct_set_unique
haftmann
parents: 35510
diff changeset
  3951
    by (simp add: sorted_distinct_set_unique)
c0db094d0d80 moved lemma map_sorted_distinct_set_unique
haftmann
parents: 35510
diff changeset
  3952
  moreover with `inj_on f (set xs \<union> set ys)` show "xs = ys"
c0db094d0d80 moved lemma map_sorted_distinct_set_unique
haftmann
parents: 35510
diff changeset
  3953
    by (blast intro: map_inj_on)
c0db094d0d80 moved lemma map_sorted_distinct_set_unique
haftmann
parents: 35510
diff changeset
  3954
qed
c0db094d0d80 moved lemma map_sorted_distinct_set_unique
haftmann
parents: 35510
diff changeset
  3955
24645
1af302128da2 Generalized [_.._] from nat to linear orders
nipkow
parents: 24640
diff changeset
  3956
lemma finite_sorted_distinct_unique:
1af302128da2 Generalized [_.._] from nat to linear orders
nipkow
parents: 24640
diff changeset
  3957
shows "finite A \<Longrightarrow> EX! xs. set xs = A & sorted xs & distinct xs"
1af302128da2 Generalized [_.._] from nat to linear orders
nipkow
parents: 24640
diff changeset
  3958
apply(drule finite_distinct_list)
1af302128da2 Generalized [_.._] from nat to linear orders
nipkow
parents: 24640
diff changeset
  3959
apply clarify
1af302128da2 Generalized [_.._] from nat to linear orders
nipkow
parents: 24640
diff changeset
  3960
apply(rule_tac a="sort xs" in ex1I)
1af302128da2 Generalized [_.._] from nat to linear orders
nipkow
parents: 24640
diff changeset
  3961
apply (auto simp: sorted_distinct_set_unique)
1af302128da2 Generalized [_.._] from nat to linear orders
nipkow
parents: 24640
diff changeset
  3962
done
1af302128da2 Generalized [_.._] from nat to linear orders
nipkow
parents: 24640
diff changeset
  3963
39915
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  3964
lemma
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  3965
  assumes "sorted xs"
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  3966
  shows sorted_take: "sorted (take n xs)"
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  3967
  and sorted_drop: "sorted (drop n xs)"
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  3968
proof -
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  3969
  from assms have "sorted (take n xs @ drop n xs)" by simp
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  3970
  then show "sorted (take n xs)" and "sorted (drop n xs)"
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  3971
    unfolding sorted_append by simp_all
29626
6f8aada233c1 sorted_take, sorted_drop
haftmann
parents: 29509
diff changeset
  3972
qed
6f8aada233c1 sorted_take, sorted_drop
haftmann
parents: 29509
diff changeset
  3973
33639
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3974
lemma sorted_dropWhile: "sorted xs \<Longrightarrow> sorted (dropWhile P xs)"
39915
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  3975
  by (auto dest: sorted_drop simp add: dropWhile_eq_drop)
33639
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3976
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33593
diff changeset
  3977
lemma sorted_takeWhile: "sorted xs \<Longrightarrow> sorted (takeWhile P xs)"
39915
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39774
diff changeset
  3978
  by (subst takeWhile_eq_take) (auto dest: sorted_take)
29626
6f8aada233c1 sorted_take, sorted_drop
haftmann
parents: 29509
diff changeset
  3979
34933
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3980
lemma sorted_filter:
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3981
  "sorted (map f xs) \<Longrightarrow> sorted (map f (filter P xs))"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3982
  by (induct xs) (simp_all add: sorted_Cons)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3983
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3984
lemma foldr_max_sorted:
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3985
  assumes "sorted (rev xs)"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3986
  shows "foldr max xs y = (if xs = [] then y else max (xs ! 0) y)"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3987
using assms proof (induct xs)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3988
  case (Cons x xs)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3989
  moreover hence "sorted (rev xs)" using sorted_append by auto
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3990
  ultimately show ?case
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3991
    by (cases xs, auto simp add: sorted_append max_def)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3992
qed simp
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3993
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3994
lemma filter_equals_takeWhile_sorted_rev:
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3995
  assumes sorted: "sorted (rev (map f xs))"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3996
  shows "[x \<leftarrow> xs. t < f x] = takeWhile (\<lambda> x. t < f x) xs"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3997
    (is "filter ?P xs = ?tW")
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3998
proof (rule takeWhile_eq_filter[symmetric])
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  3999
  let "?dW" = "dropWhile ?P xs"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4000
  fix x assume "x \<in> set ?dW"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4001
  then obtain i where i: "i < length ?dW" and nth_i: "x = ?dW ! i"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4002
    unfolding in_set_conv_nth by auto
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4003
  hence "length ?tW + i < length (?tW @ ?dW)"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4004
    unfolding length_append by simp
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4005
  hence i': "length (map f ?tW) + i < length (map f xs)" by simp
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4006
  have "(map f ?tW @ map f ?dW) ! (length (map f ?tW) + i) \<le>
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4007
        (map f ?tW @ map f ?dW) ! (length (map f ?tW) + 0)"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4008
    using sorted_rev_nth_mono[OF sorted _ i', of "length ?tW"]
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4009
    unfolding map_append[symmetric] by simp
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4010
  hence "f x \<le> f (?dW ! 0)"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4011
    unfolding nth_append_length_plus nth_i
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4012
    using i preorder_class.le_less_trans[OF le0 i] by simp
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4013
  also have "... \<le> t"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4014
    using hd_dropWhile[of "?P" xs] le0[THEN preorder_class.le_less_trans, OF i]
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4015
    using hd_conv_nth[of "?dW"] by simp
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4016
  finally show "\<not> t < f x" by simp
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4017
qed
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4018
40210
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  4019
lemma insort_insert_key_triv:
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  4020
  "f x \<in> f ` set xs \<Longrightarrow> insort_insert_key f x xs = xs"
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  4021
  by (simp add: insort_insert_key_def)
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  4022
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  4023
lemma insort_insert_triv:
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  4024
  "x \<in> set xs \<Longrightarrow> insort_insert x xs = xs"
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  4025
  using insort_insert_key_triv [of "\<lambda>x. x"] by simp
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  4026
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  4027
lemma insort_insert_insort_key:
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  4028
  "f x \<notin> f ` set xs \<Longrightarrow> insort_insert_key f x xs = insort_key f x xs"
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  4029
  by (simp add: insort_insert_key_def)
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  4030
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  4031
lemma insort_insert_insort:
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  4032
  "x \<notin> set xs \<Longrightarrow> insort_insert x xs = insort x xs"
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  4033
  using insort_insert_insort_key [of "\<lambda>x. x"] by simp
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  4034
35608
db4045d1406e added insort_insert
haftmann
parents: 35603
diff changeset
  4035
lemma set_insort_insert:
db4045d1406e added insort_insert
haftmann
parents: 35603
diff changeset
  4036
  "set (insort_insert x xs) = insert x (set xs)"
40210
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  4037
  by (auto simp add: insort_insert_key_def set_insort)
35608
db4045d1406e added insort_insert
haftmann
parents: 35603
diff changeset
  4038
db4045d1406e added insort_insert
haftmann
parents: 35603
diff changeset
  4039
lemma distinct_insort_insert:
db4045d1406e added insort_insert
haftmann
parents: 35603
diff changeset
  4040
  assumes "distinct xs"
40210
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  4041
  shows "distinct (insort_insert_key f x xs)"
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  4042
  using assms by (induct xs) (auto simp add: insort_insert_key_def set_insort)
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  4043
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  4044
lemma sorted_insort_insert_key:
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  4045
  assumes "sorted (map f xs)"
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  4046
  shows "sorted (map f (insort_insert_key f x xs))"
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  4047
  using assms by (simp add: insort_insert_key_def sorted_insort_key)
35608
db4045d1406e added insort_insert
haftmann
parents: 35603
diff changeset
  4048
db4045d1406e added insort_insert
haftmann
parents: 35603
diff changeset
  4049
lemma sorted_insort_insert:
db4045d1406e added insort_insert
haftmann
parents: 35603
diff changeset
  4050
  assumes "sorted xs"
db4045d1406e added insort_insert
haftmann
parents: 35603
diff changeset
  4051
  shows "sorted (insort_insert x xs)"
40210
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  4052
  using assms sorted_insort_insert_key [of "\<lambda>x. x"] by simp
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  4053
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  4054
lemma filter_insort_triv:
37107
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4055
  "\<not> P x \<Longrightarrow> filter P (insort_key f x xs) = filter P xs"
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4056
  by (induct xs) simp_all
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4057
40210
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  4058
lemma filter_insort:
37107
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4059
  "sorted (map f xs) \<Longrightarrow> P x \<Longrightarrow> filter P (insort_key f x xs) = insort_key f x (filter P xs)"
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4060
  using assms by (induct xs)
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4061
    (auto simp add: sorted_Cons, subst insort_is_Cons, auto)
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4062
40210
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  4063
lemma filter_sort:
37107
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4064
  "filter P (sort_key f xs) = sort_key f (filter P xs)"
40210
aee7ef725330 sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents: 40195
diff changeset
  4065
  by (induct xs) (simp_all add: filter_insort_triv filter_insort)
37107
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4066
40304
62bdd1bfcd90 lemmas sorted_map_same, sorted_same
haftmann
parents: 40230
diff changeset
  4067
lemma sorted_map_same:
62bdd1bfcd90 lemmas sorted_map_same, sorted_same
haftmann
parents: 40230
diff changeset
  4068
  "sorted (map f [x\<leftarrow>xs. f x = g xs])"
62bdd1bfcd90 lemmas sorted_map_same, sorted_same
haftmann
parents: 40230
diff changeset
  4069
proof (induct xs arbitrary: g)
37107
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4070
  case Nil then show ?case by simp
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4071
next
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4072
  case (Cons x xs)
40304
62bdd1bfcd90 lemmas sorted_map_same, sorted_same
haftmann
parents: 40230
diff changeset
  4073
  then have "sorted (map f [y\<leftarrow>xs . f y = (\<lambda>xs. f x) xs])" .
62bdd1bfcd90 lemmas sorted_map_same, sorted_same
haftmann
parents: 40230
diff changeset
  4074
  moreover from Cons have "sorted (map f [y\<leftarrow>xs . f y = (g \<circ> Cons x) xs])" .
37107
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4075
  ultimately show ?case by (simp_all add: sorted_Cons)
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4076
qed
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4077
40304
62bdd1bfcd90 lemmas sorted_map_same, sorted_same
haftmann
parents: 40230
diff changeset
  4078
lemma sorted_same:
62bdd1bfcd90 lemmas sorted_map_same, sorted_same
haftmann
parents: 40230
diff changeset
  4079
  "sorted [x\<leftarrow>xs. x = g xs]"
62bdd1bfcd90 lemmas sorted_map_same, sorted_same
haftmann
parents: 40230
diff changeset
  4080
  using sorted_map_same [of "\<lambda>x. x"] by simp
62bdd1bfcd90 lemmas sorted_map_same, sorted_same
haftmann
parents: 40230
diff changeset
  4081
37107
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4082
lemma remove1_insort [simp]:
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4083
  "remove1 x (insort x xs) = xs"
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4084
  by (induct xs) simp_all
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4085
24616
fac3dd4ade83 sorting
nipkow
parents: 24566
diff changeset
  4086
end
fac3dd4ade83 sorting
nipkow
parents: 24566
diff changeset
  4087
25277
95128fcdd7e8 added lemmas
nipkow
parents: 25221
diff changeset
  4088
lemma sorted_upt[simp]: "sorted[i..<j]"
95128fcdd7e8 added lemmas
nipkow
parents: 25221
diff changeset
  4089
by (induct j) (simp_all add:sorted_append)
95128fcdd7e8 added lemmas
nipkow
parents: 25221
diff changeset
  4090
32415
1dddf2f64266 got rid of complicated class finite_intvl_succ and defined "upto" directly on int, the only instance of the class.
nipkow
parents: 32078
diff changeset
  4091
lemma sorted_upto[simp]: "sorted[i..j]"
1dddf2f64266 got rid of complicated class finite_intvl_succ and defined "upto" directly on int, the only instance of the class.
nipkow
parents: 32078
diff changeset
  4092
apply(induct i j rule:upto.induct)
1dddf2f64266 got rid of complicated class finite_intvl_succ and defined "upto" directly on int, the only instance of the class.
nipkow
parents: 32078
diff changeset
  4093
apply(subst upto.simps)
1dddf2f64266 got rid of complicated class finite_intvl_succ and defined "upto" directly on int, the only instance of the class.
nipkow
parents: 32078
diff changeset
  4094
apply(simp add:sorted_Cons)
1dddf2f64266 got rid of complicated class finite_intvl_succ and defined "upto" directly on int, the only instance of the class.
nipkow
parents: 32078
diff changeset
  4095
done
1dddf2f64266 got rid of complicated class finite_intvl_succ and defined "upto" directly on int, the only instance of the class.
nipkow
parents: 32078
diff changeset
  4096
35115
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
  4097
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
  4098
subsubsection {* @{const transpose} on sorted lists *}
34933
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4099
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4100
lemma sorted_transpose[simp]:
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4101
  shows "sorted (rev (map length (transpose xs)))"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4102
  by (auto simp: sorted_equals_nth_mono rev_nth nth_transpose
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4103
    length_filter_conv_card intro: card_mono)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4104
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4105
lemma transpose_max_length:
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4106
  "foldr (\<lambda>xs. max (length xs)) (transpose xs) 0 = length [x \<leftarrow> xs. x \<noteq> []]"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4107
  (is "?L = ?R")
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4108
proof (cases "transpose xs = []")
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4109
  case False
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4110
  have "?L = foldr max (map length (transpose xs)) 0"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4111
    by (simp add: foldr_map comp_def)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4112
  also have "... = length (transpose xs ! 0)"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4113
    using False sorted_transpose by (simp add: foldr_max_sorted)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4114
  finally show ?thesis
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4115
    using False by (simp add: nth_transpose)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4116
next
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4117
  case True
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4118
  hence "[x \<leftarrow> xs. x \<noteq> []] = []"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4119
    by (auto intro!: filter_False simp: transpose_empty)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4120
  thus ?thesis by (simp add: transpose_empty True)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4121
qed
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4122
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4123
lemma length_transpose_sorted:
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4124
  fixes xs :: "'a list list"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4125
  assumes sorted: "sorted (rev (map length xs))"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4126
  shows "length (transpose xs) = (if xs = [] then 0 else length (xs ! 0))"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4127
proof (cases "xs = []")
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4128
  case False
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4129
  thus ?thesis
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4130
    using foldr_max_sorted[OF sorted] False
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4131
    unfolding length_transpose foldr_map comp_def
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4132
    by simp
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4133
qed simp
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4134
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4135
lemma nth_nth_transpose_sorted[simp]:
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4136
  fixes xs :: "'a list list"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4137
  assumes sorted: "sorted (rev (map length xs))"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4138
  and i: "i < length (transpose xs)"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4139
  and j: "j < length [ys \<leftarrow> xs. i < length ys]"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4140
  shows "transpose xs ! i ! j = xs ! j  ! i"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4141
  using j filter_equals_takeWhile_sorted_rev[OF sorted, of i]
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4142
    nth_transpose[OF i] nth_map[OF j]
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4143
  by (simp add: takeWhile_nth)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4144
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4145
lemma transpose_column_length:
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4146
  fixes xs :: "'a list list"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4147
  assumes sorted: "sorted (rev (map length xs))" and "i < length xs"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4148
  shows "length (filter (\<lambda>ys. i < length ys) (transpose xs)) = length (xs ! i)"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4149
proof -
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4150
  have "xs \<noteq> []" using `i < length xs` by auto
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4151
  note filter_equals_takeWhile_sorted_rev[OF sorted, simp]
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4152
  { fix j assume "j \<le> i"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4153
    note sorted_rev_nth_mono[OF sorted, of j i, simplified, OF this `i < length xs`]
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4154
  } note sortedE = this[consumes 1]
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4155
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4156
  have "{j. j < length (transpose xs) \<and> i < length (transpose xs ! j)}
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4157
    = {..< length (xs ! i)}"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4158
  proof safe
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4159
    fix j
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4160
    assume "j < length (transpose xs)" and "i < length (transpose xs ! j)"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4161
    with this(2) nth_transpose[OF this(1)]
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4162
    have "i < length (takeWhile (\<lambda>ys. j < length ys) xs)" by simp
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4163
    from nth_mem[OF this] takeWhile_nth[OF this]
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4164
    show "j < length (xs ! i)" by (auto dest: set_takeWhileD)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4165
  next
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4166
    fix j assume "j < length (xs ! i)"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4167
    thus "j < length (transpose xs)"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4168
      using foldr_max_sorted[OF sorted] `xs \<noteq> []` sortedE[OF le0]
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4169
      by (auto simp: length_transpose comp_def foldr_map)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4170
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4171
    have "Suc i \<le> length (takeWhile (\<lambda>ys. j < length ys) xs)"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4172
      using `i < length xs` `j < length (xs ! i)` less_Suc_eq_le
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4173
      by (auto intro!: length_takeWhile_less_P_nth dest!: sortedE)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4174
    with nth_transpose[OF `j < length (transpose xs)`]
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4175
    show "i < length (transpose xs ! j)" by simp
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4176
  qed
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4177
  thus ?thesis by (simp add: length_filter_conv_card)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4178
qed
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4179
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4180
lemma transpose_column:
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4181
  fixes xs :: "'a list list"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4182
  assumes sorted: "sorted (rev (map length xs))" and "i < length xs"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4183
  shows "map (\<lambda>ys. ys ! i) (filter (\<lambda>ys. i < length ys) (transpose xs))
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4184
    = xs ! i" (is "?R = _")
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4185
proof (rule nth_equalityI, safe)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4186
  show length: "length ?R = length (xs ! i)"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4187
    using transpose_column_length[OF assms] by simp
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4188
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4189
  fix j assume j: "j < length ?R"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4190
  note * = less_le_trans[OF this, unfolded length_map, OF length_filter_le]
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4191
  from j have j_less: "j < length (xs ! i)" using length by simp
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4192
  have i_less_tW: "Suc i \<le> length (takeWhile (\<lambda>ys. Suc j \<le> length ys) xs)"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4193
  proof (rule length_takeWhile_less_P_nth)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4194
    show "Suc i \<le> length xs" using `i < length xs` by simp
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4195
    fix k assume "k < Suc i"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4196
    hence "k \<le> i" by auto
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4197
    with sorted_rev_nth_mono[OF sorted this] `i < length xs`
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4198
    have "length (xs ! i) \<le> length (xs ! k)" by simp
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4199
    thus "Suc j \<le> length (xs ! k)" using j_less by simp
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4200
  qed
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4201
  have i_less_filter: "i < length [ys\<leftarrow>xs . j < length ys]"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4202
    unfolding filter_equals_takeWhile_sorted_rev[OF sorted, of j]
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4203
    using i_less_tW by (simp_all add: Suc_le_eq)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4204
  from j show "?R ! j = xs ! i ! j"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4205
    unfolding filter_equals_takeWhile_sorted_rev[OF sorted_transpose, of i]
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4206
    by (simp add: takeWhile_nth nth_nth_transpose_sorted[OF sorted * i_less_filter])
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4207
qed
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4208
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4209
lemma transpose_transpose:
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4210
  fixes xs :: "'a list list"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4211
  assumes sorted: "sorted (rev (map length xs))"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4212
  shows "transpose (transpose xs) = takeWhile (\<lambda>x. x \<noteq> []) xs" (is "?L = ?R")
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4213
proof -
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4214
  have len: "length ?L = length ?R"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4215
    unfolding length_transpose transpose_max_length
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4216
    using filter_equals_takeWhile_sorted_rev[OF sorted, of 0]
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4217
    by simp
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4218
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4219
  { fix i assume "i < length ?R"
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4220
    with less_le_trans[OF _ length_takeWhile_le[of _ xs]]
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4221
    have "i < length xs" by simp
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4222
  } note * = this
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4223
  show ?thesis
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4224
    by (rule nth_equalityI)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4225
       (simp_all add: len nth_transpose transpose_column[OF sorted] * takeWhile_nth)
0652d00305be Add transpose to the List-theory.
hoelzl
parents: 34917
diff changeset
  4226
qed
24616
fac3dd4ade83 sorting
nipkow
parents: 24566
diff changeset
  4227
34934
440605046777 Added transpose_rectangle, when the input list is rectangular.
hoelzl
parents: 34933
diff changeset
  4228
theorem transpose_rectangle:
440605046777 Added transpose_rectangle, when the input list is rectangular.
hoelzl
parents: 34933
diff changeset
  4229
  assumes "xs = [] \<Longrightarrow> n = 0"
440605046777 Added transpose_rectangle, when the input list is rectangular.
hoelzl
parents: 34933
diff changeset
  4230
  assumes rect: "\<And> i. i < length xs \<Longrightarrow> length (xs ! i) = n"
440605046777 Added transpose_rectangle, when the input list is rectangular.
hoelzl
parents: 34933
diff changeset
  4231
  shows "transpose xs = map (\<lambda> i. map (\<lambda> j. xs ! j ! i) [0..<length xs]) [0..<n]"
440605046777 Added transpose_rectangle, when the input list is rectangular.
hoelzl
parents: 34933
diff changeset
  4232
    (is "?trans = ?map")
440605046777 Added transpose_rectangle, when the input list is rectangular.
hoelzl
parents: 34933
diff changeset
  4233
proof (rule nth_equalityI)
440605046777 Added transpose_rectangle, when the input list is rectangular.
hoelzl
parents: 34933
diff changeset
  4234
  have "sorted (rev (map length xs))"
440605046777 Added transpose_rectangle, when the input list is rectangular.
hoelzl
parents: 34933
diff changeset
  4235
    by (auto simp: rev_nth rect intro!: sorted_nth_monoI)
440605046777 Added transpose_rectangle, when the input list is rectangular.
hoelzl
parents: 34933
diff changeset
  4236
  from foldr_max_sorted[OF this] assms
440605046777 Added transpose_rectangle, when the input list is rectangular.
hoelzl
parents: 34933
diff changeset
  4237
  show len: "length ?trans = length ?map"
440605046777 Added transpose_rectangle, when the input list is rectangular.
hoelzl
parents: 34933
diff changeset
  4238
    by (simp_all add: length_transpose foldr_map comp_def)
440605046777 Added transpose_rectangle, when the input list is rectangular.
hoelzl
parents: 34933
diff changeset
  4239
  moreover
440605046777 Added transpose_rectangle, when the input list is rectangular.
hoelzl
parents: 34933
diff changeset
  4240
  { fix i assume "i < n" hence "[ys\<leftarrow>xs . i < length ys] = xs"
440605046777 Added transpose_rectangle, when the input list is rectangular.
hoelzl
parents: 34933
diff changeset
  4241
      using rect by (auto simp: in_set_conv_nth intro!: filter_True) }
440605046777 Added transpose_rectangle, when the input list is rectangular.
hoelzl
parents: 34933
diff changeset
  4242
  ultimately show "\<forall>i < length ?trans. ?trans ! i = ?map ! i"
440605046777 Added transpose_rectangle, when the input list is rectangular.
hoelzl
parents: 34933
diff changeset
  4243
    by (auto simp: nth_transpose intro: nth_equalityI)
440605046777 Added transpose_rectangle, when the input list is rectangular.
hoelzl
parents: 34933
diff changeset
  4244
qed
24616
fac3dd4ade83 sorting
nipkow
parents: 24566
diff changeset
  4245
35115
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
  4246
25069
081189141a6e added sorted_list_of_set
nipkow
parents: 25062
diff changeset
  4247
subsubsection {* @{text sorted_list_of_set} *}
081189141a6e added sorted_list_of_set
nipkow
parents: 25062
diff changeset
  4248
081189141a6e added sorted_list_of_set
nipkow
parents: 25062
diff changeset
  4249
text{* This function maps (finite) linearly ordered sets to sorted
081189141a6e added sorted_list_of_set
nipkow
parents: 25062
diff changeset
  4250
lists. Warning: in most cases it is not a good idea to convert from
081189141a6e added sorted_list_of_set
nipkow
parents: 25062
diff changeset
  4251
sets to lists but one should convert in the other direction (via
081189141a6e added sorted_list_of_set
nipkow
parents: 25062
diff changeset
  4252
@{const set}). *}
081189141a6e added sorted_list_of_set
nipkow
parents: 25062
diff changeset
  4253
081189141a6e added sorted_list_of_set
nipkow
parents: 25062
diff changeset
  4254
context linorder
081189141a6e added sorted_list_of_set
nipkow
parents: 25062
diff changeset
  4255
begin
081189141a6e added sorted_list_of_set
nipkow
parents: 25062
diff changeset
  4256
35195
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  4257
definition sorted_list_of_set :: "'a set \<Rightarrow> 'a list" where
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  4258
  "sorted_list_of_set = Finite_Set.fold insort []"
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  4259
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  4260
lemma sorted_list_of_set_empty [simp]:
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  4261
  "sorted_list_of_set {} = []"
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  4262
  by (simp add: sorted_list_of_set_def)
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  4263
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  4264
lemma sorted_list_of_set_insert [simp]:
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  4265
  assumes "finite A"
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  4266
  shows "sorted_list_of_set (insert x A) = insort x (sorted_list_of_set (A - {x}))"
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  4267
proof -
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42809
diff changeset
  4268
  interpret comp_fun_commute insort by (fact comp_fun_commute_insort)
35195
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  4269
  with assms show ?thesis by (simp add: sorted_list_of_set_def fold_insert_remove)
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  4270
qed
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  4271
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  4272
lemma sorted_list_of_set [simp]:
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  4273
  "finite A \<Longrightarrow> set (sorted_list_of_set A) = A \<and> sorted (sorted_list_of_set A) 
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  4274
    \<and> distinct (sorted_list_of_set A)"
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  4275
  by (induct A rule: finite_induct) (simp_all add: set_insort sorted_insort distinct_insort)
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  4276
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  4277
lemma sorted_list_of_set_sort_remdups:
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  4278
  "sorted_list_of_set (set xs) = sort (remdups xs)"
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  4279
proof -
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42809
diff changeset
  4280
  interpret comp_fun_commute insort by (fact comp_fun_commute_insort)
35195
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  4281
  show ?thesis by (simp add: sort_foldl_insort sorted_list_of_set_def fold_set_remdups)
5163c2d00904 more lemmas about sort(_key)
haftmann
parents: 35115
diff changeset
  4282
qed
25069
081189141a6e added sorted_list_of_set
nipkow
parents: 25062
diff changeset
  4283
37107
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4284
lemma sorted_list_of_set_remove:
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4285
  assumes "finite A"
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4286
  shows "sorted_list_of_set (A - {x}) = remove1 x (sorted_list_of_set A)"
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4287
proof (cases "x \<in> A")
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4288
  case False with assms have "x \<notin> set (sorted_list_of_set A)" by simp
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4289
  with False show ?thesis by (simp add: remove1_idem)
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4290
next
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4291
  case True then obtain B where A: "A = insert x B" by (rule Set.set_insert)
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4292
  with assms show ?thesis by simp
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4293
qed
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4294
25069
081189141a6e added sorted_list_of_set
nipkow
parents: 25062
diff changeset
  4295
end
081189141a6e added sorted_list_of_set
nipkow
parents: 25062
diff changeset
  4296
37107
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4297
lemma sorted_list_of_set_range [simp]:
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4298
  "sorted_list_of_set {m..<n} = [m..<n]"
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4299
  by (rule sorted_distinct_set_unique) simp_all
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4300
1535aa1c943a more lemmas
haftmann
parents: 37020
diff changeset
  4301
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15307
diff changeset
  4302
subsubsection {* @{text lists}: the list-forming operator over sets *}
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4303
23740
d7f18c837ce7 Adapted to new package for inductive sets.
berghofe
parents: 23554
diff changeset
  4304
inductive_set
22262
96ba62dff413 Adapted to new inductive definition package.
berghofe
parents: 22143
diff changeset
  4305
  lists :: "'a set => 'a list set"
23740
d7f18c837ce7 Adapted to new package for inductive sets.
berghofe
parents: 23554
diff changeset
  4306
  for A :: "'a set"
d7f18c837ce7 Adapted to new package for inductive sets.
berghofe
parents: 23554
diff changeset
  4307
where
39613
7723505c746a more lists lemmas
nipkow
parents: 39597
diff changeset
  4308
    Nil [intro!, simp]: "[]: lists A"
7723505c746a more lists lemmas
nipkow
parents: 39597
diff changeset
  4309
  | Cons [intro!, simp, no_atp]: "[| a: A; l: lists A|] ==> a#l : lists A"
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35827
diff changeset
  4310
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35827
diff changeset
  4311
inductive_cases listsE [elim!,no_atp]: "x#l : lists A"
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35827
diff changeset
  4312
inductive_cases listspE [elim!,no_atp]: "listsp A (x # l)"
23740
d7f18c837ce7 Adapted to new package for inductive sets.
berghofe
parents: 23554
diff changeset
  4313
d7f18c837ce7 Adapted to new package for inductive sets.
berghofe
parents: 23554
diff changeset
  4314
lemma listsp_mono [mono]: "A \<le> B ==> listsp A \<le> listsp B"
34064
eee04bbbae7e avoid dependency on implicit dest rule predicate1D in proofs
haftmann
parents: 34007
diff changeset
  4315
by (rule predicate1I, erule listsp.induct, (blast dest: predicate1D)+)
26795
a27607030a1c - Explicitely applied predicate1I in a few proofs, because it is no longer
berghofe
parents: 26771
diff changeset
  4316
a27607030a1c - Explicitely applied predicate1I in a few proofs, because it is no longer
berghofe
parents: 26771
diff changeset
  4317
lemmas lists_mono = listsp_mono [to_set pred_subset_eq]
22262
96ba62dff413 Adapted to new inductive definition package.
berghofe
parents: 22143
diff changeset
  4318
22422
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 22262
diff changeset
  4319
lemma listsp_infI:
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 22262
diff changeset
  4320
  assumes l: "listsp A l" shows "listsp B l ==> listsp (inf A B) l" using l
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  4321
by induct blast+
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4322
22422
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 22262
diff changeset
  4323
lemmas lists_IntI = listsp_infI [to_set]
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 22262
diff changeset
  4324
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 22262
diff changeset
  4325
lemma listsp_inf_eq [simp]: "listsp (inf A B) = inf (listsp A) (listsp B)"
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 22262
diff changeset
  4326
proof (rule mono_inf [where f=listsp, THEN order_antisym])
22262
96ba62dff413 Adapted to new inductive definition package.
berghofe
parents: 22143
diff changeset
  4327
  show "mono listsp" by (simp add: mono_def listsp_mono)
26795
a27607030a1c - Explicitely applied predicate1I in a few proofs, because it is no longer
berghofe
parents: 26771
diff changeset
  4328
  show "inf (listsp A) (listsp B) \<le> listsp (inf A B)" by (blast intro!: listsp_infI predicate1I)
14388
04f04408b99b lemmas about card (set xs)
kleing
parents: 14343
diff changeset
  4329
qed
04f04408b99b lemmas about card (set xs)
kleing
parents: 14343
diff changeset
  4330
41075
4bed56dc95fb primitive definitions of bot/top/inf/sup for bool and fun are named with canonical suffix `_def` rather than `_eq`
haftmann
parents: 40968
diff changeset
  4331
lemmas listsp_conj_eq [simp] = listsp_inf_eq [simplified inf_fun_def inf_bool_def]
22422
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 22262
diff changeset
  4332
26795
a27607030a1c - Explicitely applied predicate1I in a few proofs, because it is no longer
berghofe
parents: 26771
diff changeset
  4333
lemmas lists_Int_eq [simp] = listsp_inf_eq [to_set pred_equals_eq]
22262
96ba62dff413 Adapted to new inductive definition package.
berghofe
parents: 22143
diff changeset
  4334
39613
7723505c746a more lists lemmas
nipkow
parents: 39597
diff changeset
  4335
lemma Cons_in_lists_iff[simp]: "x#xs : lists A \<longleftrightarrow> x:A \<and> xs : lists A"
7723505c746a more lists lemmas
nipkow
parents: 39597
diff changeset
  4336
by auto
7723505c746a more lists lemmas
nipkow
parents: 39597
diff changeset
  4337
22262
96ba62dff413 Adapted to new inductive definition package.
berghofe
parents: 22143
diff changeset
  4338
lemma append_in_listsp_conv [iff]:
96ba62dff413 Adapted to new inductive definition package.
berghofe
parents: 22143
diff changeset
  4339
     "(listsp A (xs @ ys)) = (listsp A xs \<and> listsp A ys)"
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4340
by (induct xs) auto
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4341
22262
96ba62dff413 Adapted to new inductive definition package.
berghofe
parents: 22143
diff changeset
  4342
lemmas append_in_lists_conv [iff] = append_in_listsp_conv [to_set]
96ba62dff413 Adapted to new inductive definition package.
berghofe
parents: 22143
diff changeset
  4343
96ba62dff413 Adapted to new inductive definition package.
berghofe
parents: 22143
diff changeset
  4344
lemma in_listsp_conv_set: "(listsp A xs) = (\<forall>x \<in> set xs. A x)"
96ba62dff413 Adapted to new inductive definition package.
berghofe
parents: 22143
diff changeset
  4345
-- {* eliminate @{text listsp} in favour of @{text set} *}
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4346
by (induct xs) auto
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4347
22262
96ba62dff413 Adapted to new inductive definition package.
berghofe
parents: 22143
diff changeset
  4348
lemmas in_lists_conv_set = in_listsp_conv_set [to_set]
96ba62dff413 Adapted to new inductive definition package.
berghofe
parents: 22143
diff changeset
  4349
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35827
diff changeset
  4350
lemma in_listspD [dest!,no_atp]: "listsp A xs ==> \<forall>x\<in>set xs. A x"
22262
96ba62dff413 Adapted to new inductive definition package.
berghofe
parents: 22143
diff changeset
  4351
by (rule in_listsp_conv_set [THEN iffD1])
96ba62dff413 Adapted to new inductive definition package.
berghofe
parents: 22143
diff changeset
  4352
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35827
diff changeset
  4353
lemmas in_listsD [dest!,no_atp] = in_listspD [to_set]
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35827
diff changeset
  4354
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35827
diff changeset
  4355
lemma in_listspI [intro!,no_atp]: "\<forall>x\<in>set xs. A x ==> listsp A xs"
22262
96ba62dff413 Adapted to new inductive definition package.
berghofe
parents: 22143
diff changeset
  4356
by (rule in_listsp_conv_set [THEN iffD2])
96ba62dff413 Adapted to new inductive definition package.
berghofe
parents: 22143
diff changeset
  4357
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35827
diff changeset
  4358
lemmas in_listsI [intro!,no_atp] = in_listspI [to_set]
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4359
39597
48f63a6c7f85 new lemma
nipkow
parents: 39534
diff changeset
  4360
lemma lists_eq_set: "lists A = {xs. set xs <= A}"
48f63a6c7f85 new lemma
nipkow
parents: 39534
diff changeset
  4361
by auto
48f63a6c7f85 new lemma
nipkow
parents: 39534
diff changeset
  4362
39613
7723505c746a more lists lemmas
nipkow
parents: 39597
diff changeset
  4363
lemma lists_empty [simp]: "lists {} = {[]}"
7723505c746a more lists lemmas
nipkow
parents: 39597
diff changeset
  4364
by auto
7723505c746a more lists lemmas
nipkow
parents: 39597
diff changeset
  4365
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4366
lemma lists_UNIV [simp]: "lists UNIV = UNIV"
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4367
by auto
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4368
17086
0eb0c9259dd7 added quite a few functions for code generation
nipkow
parents: 16998
diff changeset
  4369
35115
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
  4370
subsubsection {* Inductive definition for membership *}
17086
0eb0c9259dd7 added quite a few functions for code generation
nipkow
parents: 16998
diff changeset
  4371
23740
d7f18c837ce7 Adapted to new package for inductive sets.
berghofe
parents: 23554
diff changeset
  4372
inductive ListMem :: "'a \<Rightarrow> 'a list \<Rightarrow> bool"
22262
96ba62dff413 Adapted to new inductive definition package.
berghofe
parents: 22143
diff changeset
  4373
where
96ba62dff413 Adapted to new inductive definition package.
berghofe
parents: 22143
diff changeset
  4374
    elem:  "ListMem x (x # xs)"
96ba62dff413 Adapted to new inductive definition package.
berghofe
parents: 22143
diff changeset
  4375
  | insert:  "ListMem x xs \<Longrightarrow> ListMem x (y # xs)"
96ba62dff413 Adapted to new inductive definition package.
berghofe
parents: 22143
diff changeset
  4376
96ba62dff413 Adapted to new inductive definition package.
berghofe
parents: 22143
diff changeset
  4377
lemma ListMem_iff: "(ListMem x xs) = (x \<in> set xs)"
17086
0eb0c9259dd7 added quite a few functions for code generation
nipkow
parents: 16998
diff changeset
  4378
apply (rule iffI)
0eb0c9259dd7 added quite a few functions for code generation
nipkow
parents: 16998
diff changeset
  4379
 apply (induct set: ListMem)
0eb0c9259dd7 added quite a few functions for code generation
nipkow
parents: 16998
diff changeset
  4380
  apply auto
0eb0c9259dd7 added quite a few functions for code generation
nipkow
parents: 16998
diff changeset
  4381
apply (induct xs)
0eb0c9259dd7 added quite a few functions for code generation
nipkow
parents: 16998
diff changeset
  4382
 apply (auto intro: ListMem.intros)
0eb0c9259dd7 added quite a few functions for code generation
nipkow
parents: 16998
diff changeset
  4383
done
0eb0c9259dd7 added quite a few functions for code generation
nipkow
parents: 16998
diff changeset
  4384
0eb0c9259dd7 added quite a few functions for code generation
nipkow
parents: 16998
diff changeset
  4385
35115
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
  4386
subsubsection {* Lists as Cartesian products *}
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4387
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4388
text{*@{text"set_Cons A Xs"}: the set of lists with head drawn from
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4389
@{term A} and tail drawn from @{term Xs}.*}
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4390
34941
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
  4391
definition
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
  4392
  set_Cons :: "'a set \<Rightarrow> 'a list set \<Rightarrow> 'a list set" where
37767
a2b7a20d6ea3 dropped superfluous [code del]s
haftmann
parents: 37605
diff changeset
  4393
  "set_Cons A XS = {z. \<exists>x xs. z = x # xs \<and> x \<in> A \<and> xs \<in> XS}"
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4394
17724
e969fc0a4925 simprules need names
paulson
parents: 17629
diff changeset
  4395
lemma set_Cons_sing_Nil [simp]: "set_Cons A {[]} = (%x. [x])`A"
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4396
by (auto simp add: set_Cons_def)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4397
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4398
text{*Yields the set of lists, all of the same length as the argument and
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4399
with elements drawn from the corresponding element of the argument.*}
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4400
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4401
primrec
34941
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
  4402
  listset :: "'a set list \<Rightarrow> 'a list set" where
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
  4403
     "listset [] = {[]}"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
  4404
  |  "listset (A # As) = set_Cons A (listset As)"
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4405
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4406
35115
446c5063e4fd modernized translations;
wenzelm
parents: 35028
diff changeset
  4407
subsection {* Relations on Lists *}
15656
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4408
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4409
subsubsection {* Length Lexicographic Ordering *}
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4410
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4411
text{*These orderings preserve well-foundedness: shorter lists 
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4412
  precede longer lists. These ordering are not used in dictionaries.*}
34941
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
  4413
        
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
  4414
primrec -- {*The lexicographic ordering for lists of the specified length*}
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
  4415
  lexn :: "('a \<times> 'a) set \<Rightarrow> nat \<Rightarrow> ('a list \<times> 'a list) set" where
37767
a2b7a20d6ea3 dropped superfluous [code del]s
haftmann
parents: 37605
diff changeset
  4416
    "lexn r 0 = {}"
40608
6c28ab8b8166 mapper for list type; map_pair replaces prod_fun
haftmann
parents: 40593
diff changeset
  4417
  | "lexn r (Suc n) = (map_pair (%(x, xs). x#xs) (%(x, xs). x#xs) ` (r <*lex*> lexn r n)) Int
34941
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
  4418
      {(xs, ys). length xs = Suc n \<and> length ys = Suc n}"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
  4419
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
  4420
definition
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
  4421
  lex :: "('a \<times> 'a) set \<Rightarrow> ('a list \<times> 'a list) set" where
37767
a2b7a20d6ea3 dropped superfluous [code del]s
haftmann
parents: 37605
diff changeset
  4422
  "lex r = (\<Union>n. lexn r n)" -- {*Holds only between lists of the same length*}
34941
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
  4423
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
  4424
definition
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
  4425
  lenlex :: "('a \<times> 'a) set => ('a list \<times> 'a list) set" where
37767
a2b7a20d6ea3 dropped superfluous [code del]s
haftmann
parents: 37605
diff changeset
  4426
  "lenlex r = inv_image (less_than <*lex*> lex r) (\<lambda>xs. (length xs, xs))"
34941
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
  4427
        -- {*Compares lists by their length and then lexicographically*}
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4428
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4429
lemma wf_lexn: "wf r ==> wf (lexn r n)"
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4430
apply (induct n, simp, simp)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4431
apply(rule wf_subset)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4432
 prefer 2 apply (rule Int_lower1)
40608
6c28ab8b8166 mapper for list type; map_pair replaces prod_fun
haftmann
parents: 40593
diff changeset
  4433
apply(rule wf_map_pair_image)
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4434
 prefer 2 apply (rule inj_onI, auto)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4435
done
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4436
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4437
lemma lexn_length:
24526
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  4438
  "(xs, ys) : lexn r n ==> length xs = n \<and> length ys = n"
7fa202789bf6 tuned lemma; replaced !! by arbitrary
nipkow
parents: 24476
diff changeset
  4439
by (induct n arbitrary: xs ys) auto
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4440
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4441
lemma wf_lex [intro!]: "wf r ==> wf (lex r)"
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4442
apply (unfold lex_def)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4443
apply (rule wf_UN)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4444
apply (blast intro: wf_lexn, clarify)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4445
apply (rename_tac m n)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4446
apply (subgoal_tac "m \<noteq> n")
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4447
 prefer 2 apply blast
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4448
apply (blast dest: lexn_length not_sym)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4449
done
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4450
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4451
lemma lexn_conv:
15656
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4452
  "lexn r n =
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4453
    {(xs,ys). length xs = n \<and> length ys = n \<and>
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4454
    (\<exists>xys x y xs' ys'. xs= xys @ x#xs' \<and> ys= xys @ y # ys' \<and> (x, y):r)}"
18423
d7859164447f new lemmas
nipkow
parents: 18336
diff changeset
  4455
apply (induct n, simp)
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4456
apply (simp add: image_Collect lex_prod_def, safe, blast)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4457
 apply (rule_tac x = "ab # xys" in exI, simp)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4458
apply (case_tac xys, simp_all, blast)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4459
done
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4460
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4461
lemma lex_conv:
15656
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4462
  "lex r =
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4463
    {(xs,ys). length xs = length ys \<and>
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4464
    (\<exists>xys x y xs' ys'. xs = xys @ x # xs' \<and> ys = xys @ y # ys' \<and> (x, y):r)}"
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4465
by (force simp add: lex_def lexn_conv)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4466
15693
3a67e61c6e96 tuned Map, renamed lex stuff in List.
nipkow
parents: 15656
diff changeset
  4467
lemma wf_lenlex [intro!]: "wf r ==> wf (lenlex r)"
3a67e61c6e96 tuned Map, renamed lex stuff in List.
nipkow
parents: 15656
diff changeset
  4468
by (unfold lenlex_def) blast
3a67e61c6e96 tuned Map, renamed lex stuff in List.
nipkow
parents: 15656
diff changeset
  4469
3a67e61c6e96 tuned Map, renamed lex stuff in List.
nipkow
parents: 15656
diff changeset
  4470
lemma lenlex_conv:
3a67e61c6e96 tuned Map, renamed lex stuff in List.
nipkow
parents: 15656
diff changeset
  4471
    "lenlex r = {(xs,ys). length xs < length ys |
15656
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4472
                 length xs = length ys \<and> (xs, ys) : lex r}"
30198
922f944f03b2 name changes
nipkow
parents: 30128
diff changeset
  4473
by (simp add: lenlex_def Id_on_def lex_prod_def inv_image_def)
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4474
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4475
lemma Nil_notin_lex [iff]: "([], ys) \<notin> lex r"
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4476
by (simp add: lex_conv)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4477
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4478
lemma Nil2_notin_lex [iff]: "(xs, []) \<notin> lex r"
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4479
by (simp add:lex_conv)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4480
18447
da548623916a removed or modified some instances of [iff]
paulson
parents: 18423
diff changeset
  4481
lemma Cons_in_lex [simp]:
15656
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4482
    "((x # xs, y # ys) : lex r) =
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4483
      ((x, y) : r \<and> length xs = length ys | x = y \<and> (xs, ys) : lex r)"
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4484
apply (simp add: lex_conv)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4485
apply (rule iffI)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4486
 prefer 2 apply (blast intro: Cons_eq_appendI, clarify)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4487
apply (case_tac xys, simp, simp)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4488
apply blast
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4489
done
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4490
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4491
15656
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4492
subsubsection {* Lexicographic Ordering *}
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4493
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4494
text {* Classical lexicographic ordering on lists, ie. "a" < "ab" < "b".
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4495
    This ordering does \emph{not} preserve well-foundedness.
17090
603f23d71ada small mods to code lemmas
nipkow
parents: 17086
diff changeset
  4496
     Author: N. Voelker, March 2005. *} 
15656
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4497
34941
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
  4498
definition
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 34886
diff changeset
  4499
  lexord :: "('a \<times> 'a) set \<Rightarrow> ('a list \<times> 'a list) set" where
37767
a2b7a20d6ea3 dropped superfluous [code del]s
haftmann
parents: 37605
diff changeset
  4500
  "lexord r = {(x,y ). \<exists> a v. y = x @ a # v \<or>
15656
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4501
            (\<exists> u a b v w. (a,b) \<in> r \<and> x = u @ (a # v) \<and> y = u @ (b # w))}"
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4502
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4503
lemma lexord_Nil_left[simp]:  "([],y) \<in> lexord r = (\<exists> a x. y = a # x)"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  4504
by (unfold lexord_def, induct_tac y, auto) 
15656
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4505
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4506
lemma lexord_Nil_right[simp]: "(x,[]) \<notin> lexord r"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  4507
by (unfold lexord_def, induct_tac x, auto)
15656
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4508
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4509
lemma lexord_cons_cons[simp]:
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4510
     "((a # x, b # y) \<in> lexord r) = ((a,b)\<in> r | (a = b & (x,y)\<in> lexord r))"
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4511
  apply (unfold lexord_def, safe, simp_all)
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4512
  apply (case_tac u, simp, simp)
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4513
  apply (case_tac u, simp, clarsimp, blast, blast, clarsimp)
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4514
  apply (erule_tac x="b # u" in allE)
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4515
  by force
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4516
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4517
lemmas lexord_simps = lexord_Nil_left lexord_Nil_right lexord_cons_cons
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4518
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4519
lemma lexord_append_rightI: "\<exists> b z. y = b # z \<Longrightarrow> (x, x @ y) \<in> lexord r"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  4520
by (induct_tac x, auto)  
15656
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4521
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4522
lemma lexord_append_left_rightI:
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4523
     "(a,b) \<in> r \<Longrightarrow> (u @ a # x, u @ b # y) \<in> lexord r"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  4524
by (induct_tac u, auto)
15656
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4525
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4526
lemma lexord_append_leftI: " (u,v) \<in> lexord r \<Longrightarrow> (x @ u, x @ v) \<in> lexord r"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  4527
by (induct x, auto)
15656
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4528
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4529
lemma lexord_append_leftD:
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4530
     "\<lbrakk> (x @ u, x @ v) \<in> lexord r; (! a. (a,a) \<notin> r) \<rbrakk> \<Longrightarrow> (u,v) \<in> lexord r"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  4531
by (erule rev_mp, induct_tac x, auto)
15656
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4532
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4533
lemma lexord_take_index_conv: 
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4534
   "((x,y) : lexord r) = 
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4535
    ((length x < length y \<and> take (length x) y = x) \<or> 
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4536
     (\<exists>i. i < min(length x)(length y) & take i x = take i y & (x!i,y!i) \<in> r))"
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4537
  apply (unfold lexord_def Let_def, clarsimp) 
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4538
  apply (rule_tac f = "(% a b. a \<or> b)" in arg_cong2)
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4539
  apply auto 
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4540
  apply (rule_tac x="hd (drop (length x) y)" in exI)
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4541
  apply (rule_tac x="tl (drop (length x) y)" in exI)
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4542
  apply (erule subst, simp add: min_def) 
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4543
  apply (rule_tac x ="length u" in exI, simp) 
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4544
  apply (rule_tac x ="take i x" in exI) 
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4545
  apply (rule_tac x ="x ! i" in exI) 
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4546
  apply (rule_tac x ="y ! i" in exI, safe) 
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4547
  apply (rule_tac x="drop (Suc i) x" in exI)
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4548
  apply (drule sym, simp add: drop_Suc_conv_tl) 
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4549
  apply (rule_tac x="drop (Suc i) y" in exI)
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4550
  by (simp add: drop_Suc_conv_tl) 
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4551
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4552
-- {* lexord is extension of partial ordering List.lex *} 
41986
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4553
lemma lexord_lex: "(x,y) \<in> lex r = ((x,y) \<in> lexord r \<and> length x = length y)"
15656
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4554
  apply (rule_tac x = y in spec) 
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4555
  apply (induct_tac x, clarsimp) 
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4556
  by (clarify, case_tac x, simp, force)
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4557
41986
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4558
lemma lexord_irreflexive: "ALL x. (x,x) \<notin> r \<Longrightarrow> (xs,xs) \<notin> lexord r"
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4559
by (induct xs) auto
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4560
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4561
text{* By Ren\'e Thiemann: *}
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4562
lemma lexord_partial_trans: 
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4563
  "(\<And>x y z. x \<in> set xs \<Longrightarrow> (x,y) \<in> r \<Longrightarrow> (y,z) \<in> r \<Longrightarrow> (x,z) \<in> r)
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4564
   \<Longrightarrow>  (xs,ys) \<in> lexord r  \<Longrightarrow>  (ys,zs) \<in> lexord r \<Longrightarrow>  (xs,zs) \<in> lexord r"
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4565
proof (induct xs arbitrary: ys zs)
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4566
  case Nil
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4567
  from Nil(3) show ?case unfolding lexord_def by (cases zs, auto)
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4568
next
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4569
  case (Cons x xs yys zzs)
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4570
  from Cons(3) obtain y ys where yys: "yys = y # ys" unfolding lexord_def
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4571
    by (cases yys, auto)
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4572
  note Cons = Cons[unfolded yys]
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4573
  from Cons(3) have one: "(x,y) \<in> r \<or> x = y \<and> (xs,ys) \<in> lexord r" by auto
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4574
  from Cons(4) obtain z zs where zzs: "zzs = z # zs" unfolding lexord_def
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4575
    by (cases zzs, auto)
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4576
  note Cons = Cons[unfolded zzs]
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4577
  from Cons(4) have two: "(y,z) \<in> r \<or> y = z \<and> (ys,zs) \<in> lexord r" by auto
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4578
  {
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4579
    assume "(xs,ys) \<in> lexord r" and "(ys,zs) \<in> lexord r"
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4580
    from Cons(1)[OF _ this] Cons(2)
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4581
    have "(xs,zs) \<in> lexord r" by auto
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4582
  } note ind1 = this
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4583
  {
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4584
    assume "(x,y) \<in> r" and "(y,z) \<in> r"
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4585
    from Cons(2)[OF _ this] have "(x,z) \<in> r" by auto
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4586
  } note ind2 = this
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4587
  from one two ind1 ind2
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4588
  have "(x,z) \<in> r \<or> x = z \<and> (xs,zs) \<in> lexord r" by blast
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4589
  thus ?case unfolding zzs by auto
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4590
qed
15656
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4591
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4592
lemma lexord_trans: 
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4593
    "\<lbrakk> (x, y) \<in> lexord r; (y, z) \<in> lexord r; trans r \<rbrakk> \<Longrightarrow> (x, z) \<in> lexord r"
41986
95a67e3f29ad added lemma
nipkow
parents: 41842
diff changeset
  4594
by(auto simp: trans_def intro:lexord_partial_trans)
15656
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4595
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4596
lemma lexord_transI:  "trans r \<Longrightarrow> trans (lexord r)"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  4597
by (rule transI, drule lexord_trans, blast) 
15656
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4598
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4599
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"
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4600
  apply (rule_tac x = y in spec) 
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4601
  apply (induct_tac x, rule allI) 
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4602
  apply (case_tac x, simp, simp) 
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4603
  apply (rule allI, case_tac x, simp, simp) 
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4604
  by blast
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4605
988f91b9c4ef lexicographic order by Norbert Voelker
paulson
parents: 15570
diff changeset
  4606
40230
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4607
subsubsection {* Lexicographic combination of measure functions *}
21103
367b4ad7c7cc Added "measures" combinator for lexicographic combinations of multiple measures.
krauss
parents: 21079
diff changeset
  4608
367b4ad7c7cc Added "measures" combinator for lexicographic combinations of multiple measures.
krauss
parents: 21079
diff changeset
  4609
text {* These are useful for termination proofs *}
367b4ad7c7cc Added "measures" combinator for lexicographic combinations of multiple measures.
krauss
parents: 21079
diff changeset
  4610
367b4ad7c7cc Added "measures" combinator for lexicographic combinations of multiple measures.
krauss
parents: 21079
diff changeset
  4611
definition
367b4ad7c7cc Added "measures" combinator for lexicographic combinations of multiple measures.
krauss
parents: 21079
diff changeset
  4612
  "measures fs = inv_image (lex less_than) (%a. map (%f. f a) fs)"
367b4ad7c7cc Added "measures" combinator for lexicographic combinations of multiple measures.
krauss
parents: 21079
diff changeset
  4613
44013
5cfc1c36ae97 moved recdef package to HOL/Library/Old_Recdef.thy
krauss
parents: 43594
diff changeset
  4614
lemma wf_measures[simp]: "wf (measures fs)"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  4615
unfolding measures_def
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  4616
by blast
21103
367b4ad7c7cc Added "measures" combinator for lexicographic combinations of multiple measures.
krauss
parents: 21079
diff changeset
  4617
367b4ad7c7cc Added "measures" combinator for lexicographic combinations of multiple measures.
krauss
parents: 21079
diff changeset
  4618
lemma in_measures[simp]: 
367b4ad7c7cc Added "measures" combinator for lexicographic combinations of multiple measures.
krauss
parents: 21079
diff changeset
  4619
  "(x, y) \<in> measures [] = False"
367b4ad7c7cc Added "measures" combinator for lexicographic combinations of multiple measures.
krauss
parents: 21079
diff changeset
  4620
  "(x, y) \<in> measures (f # fs)
367b4ad7c7cc Added "measures" combinator for lexicographic combinations of multiple measures.
krauss
parents: 21079
diff changeset
  4621
         = (f x < f y \<or> (f x = f y \<and> (x, y) \<in> measures fs))"  
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  4622
unfolding measures_def
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  4623
by auto
21103
367b4ad7c7cc Added "measures" combinator for lexicographic combinations of multiple measures.
krauss
parents: 21079
diff changeset
  4624
367b4ad7c7cc Added "measures" combinator for lexicographic combinations of multiple measures.
krauss
parents: 21079
diff changeset
  4625
lemma measures_less: "f x < f y ==> (x, y) \<in> measures (f#fs)"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  4626
by simp
21103
367b4ad7c7cc Added "measures" combinator for lexicographic combinations of multiple measures.
krauss
parents: 21079
diff changeset
  4627
367b4ad7c7cc Added "measures" combinator for lexicographic combinations of multiple measures.
krauss
parents: 21079
diff changeset
  4628
lemma measures_lesseq: "f x <= f y ==> (x, y) \<in> measures fs ==> (x, y) \<in> measures (f#fs)"
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24335
diff changeset
  4629
by auto
21103
367b4ad7c7cc Added "measures" combinator for lexicographic combinations of multiple measures.
krauss
parents: 21079
diff changeset
  4630
367b4ad7c7cc Added "measures" combinator for lexicographic combinations of multiple measures.
krauss
parents: 21079
diff changeset
  4631
40230
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4632
subsubsection {* Lifting Relations to Lists: one element *}
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4633
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4634
definition listrel1 :: "('a \<times> 'a) set \<Rightarrow> ('a list \<times> 'a list) set" where
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4635
"listrel1 r = {(xs,ys).
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4636
   \<exists>us z z' vs. xs = us @ z # vs \<and> (z,z') \<in> r \<and> ys = us @ z' # vs}"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4637
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4638
lemma listrel1I:
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4639
  "\<lbrakk> (x, y) \<in> r;  xs = us @ x # vs;  ys = us @ y # vs \<rbrakk> \<Longrightarrow>
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4640
  (xs, ys) \<in> listrel1 r"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4641
unfolding listrel1_def by auto
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4642
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4643
lemma listrel1E:
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4644
  "\<lbrakk> (xs, ys) \<in> listrel1 r;
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4645
     !!x y us vs. \<lbrakk> (x, y) \<in> r;  xs = us @ x # vs;  ys = us @ y # vs \<rbrakk> \<Longrightarrow> P
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4646
   \<rbrakk> \<Longrightarrow> P"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4647
unfolding listrel1_def by auto
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4648
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4649
lemma not_Nil_listrel1 [iff]: "([], xs) \<notin> listrel1 r"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4650
unfolding listrel1_def by blast
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4651
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4652
lemma not_listrel1_Nil [iff]: "(xs, []) \<notin> listrel1 r"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4653
unfolding listrel1_def by blast
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4654
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4655
lemma Cons_listrel1_Cons [iff]:
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4656
  "(x # xs, y # ys) \<in> listrel1 r \<longleftrightarrow>
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4657
   (x,y) \<in> r \<and> xs = ys \<or> x = y \<and> (xs, ys) \<in> listrel1 r"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4658
by (simp add: listrel1_def Cons_eq_append_conv) (blast)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4659
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4660
lemma listrel1I1: "(x,y) \<in> r \<Longrightarrow> (x # xs, y # xs) \<in> listrel1 r"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4661
by (metis Cons_listrel1_Cons)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4662
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4663
lemma listrel1I2: "(xs, ys) \<in> listrel1 r \<Longrightarrow> (x # xs, x # ys) \<in> listrel1 r"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4664
by (metis Cons_listrel1_Cons)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4665
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4666
lemma append_listrel1I:
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4667
  "(xs, ys) \<in> listrel1 r \<and> us = vs \<or> xs = ys \<and> (us, vs) \<in> listrel1 r
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4668
    \<Longrightarrow> (xs @ us, ys @ vs) \<in> listrel1 r"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4669
unfolding listrel1_def
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4670
by auto (blast intro: append_eq_appendI)+
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4671
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4672
lemma Cons_listrel1E1[elim!]:
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4673
  assumes "(x # xs, ys) \<in> listrel1 r"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4674
    and "\<And>y. ys = y # xs \<Longrightarrow> (x, y) \<in> r \<Longrightarrow> R"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4675
    and "\<And>zs. ys = x # zs \<Longrightarrow> (xs, zs) \<in> listrel1 r \<Longrightarrow> R"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4676
  shows R
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4677
using assms by (cases ys) blast+
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4678
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4679
lemma Cons_listrel1E2[elim!]:
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4680
  assumes "(xs, y # ys) \<in> listrel1 r"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4681
    and "\<And>x. xs = x # ys \<Longrightarrow> (x, y) \<in> r \<Longrightarrow> R"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4682
    and "\<And>zs. xs = y # zs \<Longrightarrow> (zs, ys) \<in> listrel1 r \<Longrightarrow> R"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4683
  shows R
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4684
using assms by (cases xs) blast+
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4685
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4686
lemma snoc_listrel1_snoc_iff:
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4687
  "(xs @ [x], ys @ [y]) \<in> listrel1 r
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4688
    \<longleftrightarrow> (xs, ys) \<in> listrel1 r \<and> x = y \<or> xs = ys \<and> (x,y) \<in> r" (is "?L \<longleftrightarrow> ?R")
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4689
proof
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4690
  assume ?L thus ?R
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44635
diff changeset
  4691
    by (fastforce simp: listrel1_def snoc_eq_iff_butlast butlast_append)
40230
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4692
next
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4693
  assume ?R then show ?L unfolding listrel1_def by force
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4694
qed
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4695
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4696
lemma listrel1_eq_len: "(xs,ys) \<in> listrel1 r \<Longrightarrow> length xs = length ys"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4697
unfolding listrel1_def by auto
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4698
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4699
lemma listrel1_mono:
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4700
  "r \<subseteq> s \<Longrightarrow> listrel1 r \<subseteq> listrel1 s"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4701
unfolding listrel1_def by blast
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4702
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4703
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4704
lemma listrel1_converse: "listrel1 (r^-1) = (listrel1 r)^-1"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4705
unfolding listrel1_def by blast
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4706
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4707
lemma in_listrel1_converse:
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4708
  "(x,y) : listrel1 (r^-1) \<longleftrightarrow> (x,y) : (listrel1 r)^-1"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4709
unfolding listrel1_def by blast
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4710
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4711
lemma listrel1_iff_update:
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4712
  "(xs,ys) \<in> (listrel1 r)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4713
   \<longleftrightarrow> (\<exists>y n. (xs ! n, y) \<in> r \<and> n < length xs \<and> ys = xs[n:=y])" (is "?L \<longleftrightarrow> ?R")
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4714
proof
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4715
  assume "?L"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4716
  then obtain x y u v where "xs = u @ x # v"  "ys = u @ y # v"  "(x,y) \<in> r"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4717
    unfolding listrel1_def by auto
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4718
  then have "ys = xs[length u := y]" and "length u < length xs"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4719
    and "(xs ! length u, y) \<in> r" by auto
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4720
  then show "?R" by auto
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4721
next
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4722
  assume "?R"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4723
  then obtain x y n where "(xs!n, y) \<in> r" "n < size xs" "ys = xs[n:=y]" "x = xs!n"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4724
    by auto
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4725
  then obtain u v where "xs = u @ x # v" and "ys = u @ y # v" and "(x, y) \<in> r"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4726
    by (auto intro: upd_conv_take_nth_drop id_take_nth_drop)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4727
  then show "?L" by (auto simp: listrel1_def)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4728
qed
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4729
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4730
44510
5e580115dfcd added lemma
nipkow
parents: 44013
diff changeset
  4731
text{* Accessible part and wellfoundedness: *}
40230
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4732
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4733
lemma Cons_acc_listrel1I [intro!]:
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4734
  "x \<in> acc r \<Longrightarrow> xs \<in> acc (listrel1 r) \<Longrightarrow> (x # xs) \<in> acc (listrel1 r)"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4735
apply (induct arbitrary: xs set: acc)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4736
apply (erule thin_rl)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4737
apply (erule acc_induct)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4738
apply (rule accI)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4739
apply (blast)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4740
done
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4741
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4742
lemma lists_accD: "xs \<in> lists (acc r) \<Longrightarrow> xs \<in> acc (listrel1 r)"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4743
apply (induct set: lists)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4744
 apply (rule accI)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4745
 apply simp
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4746
apply (rule accI)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4747
apply (fast dest: acc_downward)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4748
done
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4749
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4750
lemma lists_accI: "xs \<in> acc (listrel1 r) \<Longrightarrow> xs \<in> lists (acc r)"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4751
apply (induct set: acc)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4752
apply clarify
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4753
apply (rule accI)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44635
diff changeset
  4754
apply (fastforce dest!: in_set_conv_decomp[THEN iffD1] simp: listrel1_def)
40230
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4755
done
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4756
44510
5e580115dfcd added lemma
nipkow
parents: 44013
diff changeset
  4757
lemma wf_listrel1_iff[simp]: "wf(listrel1 r) = wf r"
5e580115dfcd added lemma
nipkow
parents: 44013
diff changeset
  4758
by(metis wf_acc_iff in_lists_conv_set lists_accI lists_accD Cons_in_lists_iff)
5e580115dfcd added lemma
nipkow
parents: 44013
diff changeset
  4759
40230
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4760
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4761
subsubsection {* Lifting Relations to Lists: all elements *}
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4762
23740
d7f18c837ce7 Adapted to new package for inductive sets.
berghofe
parents: 23554
diff changeset
  4763
inductive_set
d7f18c837ce7 Adapted to new package for inductive sets.
berghofe
parents: 23554
diff changeset
  4764
  listrel :: "('a * 'a)set => ('a list * 'a list)set"
d7f18c837ce7 Adapted to new package for inductive sets.
berghofe
parents: 23554
diff changeset
  4765
  for r :: "('a * 'a)set"
22262
96ba62dff413 Adapted to new inductive definition package.
berghofe
parents: 22143
diff changeset
  4766
where
23740
d7f18c837ce7 Adapted to new package for inductive sets.
berghofe
parents: 23554
diff changeset
  4767
    Nil:  "([],[]) \<in> listrel r"
d7f18c837ce7 Adapted to new package for inductive sets.
berghofe
parents: 23554
diff changeset
  4768
  | Cons: "[| (x,y) \<in> r; (xs,ys) \<in> listrel r |] ==> (x#xs, y#ys) \<in> listrel r"
d7f18c837ce7 Adapted to new package for inductive sets.
berghofe
parents: 23554
diff changeset
  4769
d7f18c837ce7 Adapted to new package for inductive sets.
berghofe
parents: 23554
diff changeset
  4770
inductive_cases listrel_Nil1 [elim!]: "([],xs) \<in> listrel r"
d7f18c837ce7 Adapted to new package for inductive sets.
berghofe
parents: 23554
diff changeset
  4771
inductive_cases listrel_Nil2 [elim!]: "(xs,[]) \<in> listrel r"
d7f18c837ce7 Adapted to new package for inductive sets.
berghofe
parents: 23554
diff changeset
  4772
inductive_cases listrel_Cons1 [elim!]: "(y#ys,xs) \<in> listrel r"
d7f18c837ce7 Adapted to new package for inductive sets.
berghofe
parents: 23554
diff changeset
  4773
inductive_cases listrel_Cons2 [elim!]: "(xs,y#ys) \<in> listrel r"
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4774
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4775
40230
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4776
lemma listrel_eq_len:  "(xs, ys) \<in> listrel r \<Longrightarrow> length xs = length ys"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4777
by(induct rule: listrel.induct) auto
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4778
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4779
lemma listrel_iff_zip: "(xs,ys) : listrel r \<longleftrightarrow>
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4780
  length xs = length ys & (\<forall>(x,y) \<in> set(zip xs ys). (x,y) \<in> r)" (is "?L \<longleftrightarrow> ?R")
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4781
proof
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4782
  assume ?L thus ?R by induct (auto intro: listrel_eq_len)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4783
next
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4784
  assume ?R thus ?L
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4785
    apply (clarify)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4786
    by (induct rule: list_induct2) (auto intro: listrel.intros)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4787
qed
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4788
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4789
lemma listrel_iff_nth: "(xs,ys) : listrel r \<longleftrightarrow>
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4790
  length xs = length ys & (\<forall>n < length xs. (xs!n, ys!n) \<in> r)" (is "?L \<longleftrightarrow> ?R")
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4791
by (auto simp add: all_set_conv_all_nth listrel_iff_zip)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4792
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4793
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4794
lemma listrel_mono: "r \<subseteq> s \<Longrightarrow> listrel r \<subseteq> listrel s"
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4795
apply clarify  
23740
d7f18c837ce7 Adapted to new package for inductive sets.
berghofe
parents: 23554
diff changeset
  4796
apply (erule listrel.induct)
d7f18c837ce7 Adapted to new package for inductive sets.
berghofe
parents: 23554
diff changeset
  4797
apply (blast intro: listrel.intros)+
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4798
done
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4799
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4800
lemma listrel_subset: "r \<subseteq> A \<times> A \<Longrightarrow> listrel r \<subseteq> lists A \<times> lists A"
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4801
apply clarify 
23740
d7f18c837ce7 Adapted to new package for inductive sets.
berghofe
parents: 23554
diff changeset
  4802
apply (erule listrel.induct, auto) 
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4803
done
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4804
30198
922f944f03b2 name changes
nipkow
parents: 30128
diff changeset
  4805
lemma listrel_refl_on: "refl_on A r \<Longrightarrow> refl_on (lists A) (listrel r)" 
922f944f03b2 name changes
nipkow
parents: 30128
diff changeset
  4806
apply (simp add: refl_on_def listrel_subset Ball_def)
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4807
apply (rule allI) 
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4808
apply (induct_tac x) 
23740
d7f18c837ce7 Adapted to new package for inductive sets.
berghofe
parents: 23554
diff changeset
  4809
apply (auto intro: listrel.intros)
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4810
done
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4811
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4812
lemma listrel_sym: "sym r \<Longrightarrow> sym (listrel r)" 
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4813
apply (auto simp add: sym_def)
23740
d7f18c837ce7 Adapted to new package for inductive sets.
berghofe
parents: 23554
diff changeset
  4814
apply (erule listrel.induct) 
d7f18c837ce7 Adapted to new package for inductive sets.
berghofe
parents: 23554
diff changeset
  4815
apply (blast intro: listrel.intros)+
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4816
done
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4817
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4818
lemma listrel_trans: "trans r \<Longrightarrow> trans (listrel r)" 
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4819
apply (simp add: trans_def)
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4820
apply (intro allI) 
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4821
apply (rule impI) 
23740
d7f18c837ce7 Adapted to new package for inductive sets.
berghofe
parents: 23554
diff changeset
  4822
apply (erule listrel.induct) 
d7f18c837ce7 Adapted to new package for inductive sets.
berghofe
parents: 23554
diff changeset
  4823
apply (blast intro: listrel.intros)+
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4824
done
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4825
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4826
theorem equiv_listrel: "equiv A r \<Longrightarrow> equiv (lists A) (listrel r)"
30198
922f944f03b2 name changes
nipkow
parents: 30128
diff changeset
  4827
by (simp add: equiv_def listrel_refl_on listrel_sym listrel_trans) 
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4828
40230
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4829
lemma listrel_rtrancl_refl[iff]: "(xs,xs) : listrel(r^*)"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4830
using listrel_refl_on[of UNIV, OF refl_rtrancl]
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4831
by(auto simp: refl_on_def)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4832
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4833
lemma listrel_rtrancl_trans:
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4834
  "\<lbrakk> (xs,ys) : listrel(r^*);  (ys,zs) : listrel(r^*) \<rbrakk>
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4835
  \<Longrightarrow> (xs,zs) : listrel(r^*)"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4836
by (metis listrel_trans trans_def trans_rtrancl)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4837
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4838
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4839
lemma listrel_Nil [simp]: "listrel r `` {[]} = {[]}"
23740
d7f18c837ce7 Adapted to new package for inductive sets.
berghofe
parents: 23554
diff changeset
  4840
by (blast intro: listrel.intros)
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4841
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4842
lemma listrel_Cons:
33318
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4843
     "listrel r `` {x#xs} = set_Cons (r``{x}) (listrel r `` {xs})"
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4844
by (auto simp add: set_Cons_def intro: listrel.intros)
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4845
40230
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4846
text {* Relating @{term listrel1}, @{term listrel} and closures: *}
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4847
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4848
lemma listrel1_rtrancl_subset_rtrancl_listrel1:
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4849
  "listrel1 (r^*) \<subseteq> (listrel1 r)^*"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4850
proof (rule subrelI)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4851
  fix xs ys assume 1: "(xs,ys) \<in> listrel1 (r^*)"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4852
  { fix x y us vs
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4853
    have "(x,y) : r^* \<Longrightarrow> (us @ x # vs, us @ y # vs) : (listrel1 r)^*"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4854
    proof(induct rule: rtrancl.induct)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4855
      case rtrancl_refl show ?case by simp
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4856
    next
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4857
      case rtrancl_into_rtrancl thus ?case
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4858
        by (metis listrel1I rtrancl.rtrancl_into_rtrancl)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4859
    qed }
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4860
  thus "(xs,ys) \<in> (listrel1 r)^*" using 1 by(blast elim: listrel1E)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4861
qed
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4862
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4863
lemma rtrancl_listrel1_eq_len: "(x,y) \<in> (listrel1 r)^* \<Longrightarrow> length x = length y"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4864
by (induct rule: rtrancl.induct) (auto intro: listrel1_eq_len)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4865
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4866
lemma rtrancl_listrel1_ConsI1:
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4867
  "(xs,ys) : (listrel1 r)^* \<Longrightarrow> (x#xs,x#ys) : (listrel1 r)^*"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4868
apply(induct rule: rtrancl.induct)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4869
 apply simp
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4870
by (metis listrel1I2 rtrancl.rtrancl_into_rtrancl)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4871
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4872
lemma rtrancl_listrel1_ConsI2:
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4873
  "(x,y) \<in> r^* \<Longrightarrow> (xs, ys) \<in> (listrel1 r)^*
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4874
  \<Longrightarrow> (x # xs, y # ys) \<in> (listrel1 r)^*"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4875
  by (blast intro: rtrancl_trans rtrancl_listrel1_ConsI1 
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4876
    subsetD[OF listrel1_rtrancl_subset_rtrancl_listrel1 listrel1I1])
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4877
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4878
lemma listrel1_subset_listrel:
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4879
  "r \<subseteq> r' \<Longrightarrow> refl r' \<Longrightarrow> listrel1 r \<subseteq> listrel(r')"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4880
by(auto elim!: listrel1E simp add: listrel_iff_zip set_zip refl_on_def)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4881
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4882
lemma listrel_reflcl_if_listrel1:
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4883
  "(xs,ys) : listrel1 r \<Longrightarrow> (xs,ys) : listrel(r^*)"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4884
by(erule listrel1E)(auto simp add: listrel_iff_zip set_zip)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4885
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4886
lemma listrel_rtrancl_eq_rtrancl_listrel1: "listrel (r^*) = (listrel1 r)^*"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4887
proof
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4888
  { fix x y assume "(x,y) \<in> listrel (r^*)"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4889
    then have "(x,y) \<in> (listrel1 r)^*"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4890
    by induct (auto intro: rtrancl_listrel1_ConsI2) }
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4891
  then show "listrel (r^*) \<subseteq> (listrel1 r)^*"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4892
    by (rule subrelI)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4893
next
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4894
  show "listrel (r^*) \<supseteq> (listrel1 r)^*"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4895
  proof(rule subrelI)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4896
    fix xs ys assume "(xs,ys) \<in> (listrel1 r)^*"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4897
    then show "(xs,ys) \<in> listrel (r^*)"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4898
    proof induct
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4899
      case base show ?case by(auto simp add: listrel_iff_zip set_zip)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4900
    next
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4901
      case (step ys zs)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4902
      thus ?case  by (metis listrel_reflcl_if_listrel1 listrel_rtrancl_trans)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4903
    qed
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4904
  qed
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4905
qed
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4906
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4907
lemma rtrancl_listrel1_if_listrel:
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4908
  "(xs,ys) : listrel r \<Longrightarrow> (xs,ys) : (listrel1 r)^*"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4909
by(metis listrel_rtrancl_eq_rtrancl_listrel1 subsetD[OF listrel_mono] r_into_rtrancl subsetI)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4910
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4911
lemma listrel_subset_rtrancl_listrel1: "listrel r \<subseteq> (listrel1 r)^*"
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4912
by(fast intro:rtrancl_listrel1_if_listrel)
be5c622e1de2 added lemmas about listrel(1)
nipkow
parents: 40210
diff changeset
  4913
15302
a643fcbc3468 Restructured List and added "rotate"
nipkow
parents: 15281
diff changeset
  4914
26749
397a1aeede7d * New attribute "termination_simp": Simp rules for termination proofs
krauss
parents: 26734
diff changeset
  4915
subsection {* Size function *}
397a1aeede7d * New attribute "termination_simp": Simp rules for termination proofs
krauss
parents: 26734
diff changeset
  4916
26875
e18574413bc4 Measure functions can now be declared via special rules, allowing for a
krauss
parents: 26795
diff changeset
  4917
lemma [measure_function]: "is_measure f \<Longrightarrow> is_measure (list_size f)"
e18574413bc4 Measure functions can now be declared via special rules, allowing for a
krauss
parents: 26795
diff changeset
  4918
by (rule is_measure_trivial)
e18574413bc4 Measure functions can now be declared via special rules, allowing for a
krauss
parents: 26795
diff changeset
  4919
e18574413bc4 Measure functions can now be declared via special rules, allowing for a
krauss
parents: 26795
diff changeset
  4920
lemma [measure_function]: "is_measure f \<Longrightarrow> is_measure (option_size f)"
e18574413bc4 Measure functions can now be declared via special rules, allowing for a
krauss
parents: 26795
diff changeset
  4921
by (rule is_measure_trivial)
e18574413bc4 Measure functions can now be declared via special rules, allowing for a
krauss
parents: 26795
diff changeset
  4922
e18574413bc4 Measure functions can now be declared via special rules, allowing for a
krauss
parents: 26795
diff changeset
  4923
lemma list_size_estimation[termination_simp]: 
e18574413bc4 Measure functions can now be declared via special rules, allowing for a
krauss
parents: 26795
diff changeset
  4924
  "x \<in> set xs \<Longrightarrow> y < f x \<Longrightarrow> y < list_size f xs"
26749
397a1aeede7d * New attribute "termination_simp": Simp rules for termination proofs
krauss
parents: 26734
diff changeset
  4925
by (induct xs) auto
397a1aeede7d * New attribute "termination_simp": Simp rules for termination proofs
krauss
parents: 26734
diff changeset
  4926
26875
e18574413bc4 Measure functions can now be declared via special rules, allowing for a
krauss
parents: 26795
diff changeset
  4927
lemma list_size_estimation'[termination_simp]: 
e18574413bc4 Measure functions can now be declared via special rules, allowing for a
krauss
parents: 26795
diff changeset
  4928
  "x \<in> set xs \<Longrightarrow> y \<le> f x \<Longrightarrow> y \<le> list_size f xs"
e18574413bc4 Measure functions can now be declared via special rules, allowing for a
krauss
parents: 26795
diff changeset
  4929
by (induct xs) auto
e18574413bc4 Measure functions can now be declared via special rules, allowing for a
krauss
parents: 26795
diff changeset
  4930
e18574413bc4 Measure functions can now be declared via special rules, allowing for a
krauss
parents: 26795
diff changeset
  4931
lemma list_size_map[simp]: "list_size f (map g xs) = list_size (f o g) xs"
e18574413bc4 Measure functions can now be declared via special rules, allowing for a
krauss
parents: 26795
diff changeset
  4932
by (induct xs) auto
e18574413bc4 Measure functions can now be declared via special rules, allowing for a
krauss
parents: 26795
diff changeset
  4933
44619
fd520fa2fb09 adding list_size_append (thanks to Rene Thiemann)
bulwahn
parents: 44618
diff changeset
  4934
lemma list_size_append[simp]: "list_size f (xs @ ys) = list_size f xs + list_size f ys"
fd520fa2fb09 adding list_size_append (thanks to Rene Thiemann)
bulwahn
parents: 44618
diff changeset
  4935
by (induct xs, auto)
fd520fa2fb09 adding list_size_append (thanks to Rene Thiemann)
bulwahn
parents: 44618
diff changeset
  4936
26875
e18574413bc4 Measure functions can now be declared via special rules, allowing for a
krauss
parents: 26795
diff changeset
  4937
lemma list_size_pointwise[termination_simp]: 
44618
f3635643a376 strengthening list_size_pointwise (thanks to Rene Thiemann)
bulwahn
parents: 44510
diff changeset
  4938
  "(\<And>x. x \<in> set xs \<Longrightarrow> f x \<le> g x) \<Longrightarrow> list_size f xs \<le> list_size g xs"
26875
e18574413bc4 Measure functions can now be declared via special rules, allowing for a
krauss
parents: 26795
diff changeset
  4939
by (induct xs) force+
26749
397a1aeede7d * New attribute "termination_simp": Simp rules for termination proofs
krauss
parents: 26734
diff changeset
  4940
31048
ac146fc38b51 refined HOL string theories and corresponding ML fragments
haftmann
parents: 31022
diff changeset
  4941
33318
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4942
subsection {* Transfer *}
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4943
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4944
definition
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4945
  embed_list :: "nat list \<Rightarrow> int list"
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4946
where
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4947
  "embed_list l = map int l"
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4948
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4949
definition
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4950
  nat_list :: "int list \<Rightarrow> bool"
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4951
where
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4952
  "nat_list l = nat_set (set l)"
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4953
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4954
definition
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4955
  return_list :: "int list \<Rightarrow> nat list"
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4956
where
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4957
  "return_list l = map nat l"
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4958
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4959
lemma transfer_nat_int_list_return_embed: "nat_list l \<longrightarrow>
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4960
    embed_list (return_list l) = l"
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4961
  unfolding embed_list_def return_list_def nat_list_def nat_set_def
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4962
  apply (induct l)
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4963
  apply auto
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4964
done
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4965
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4966
lemma transfer_nat_int_list_functions:
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4967
  "l @ m = return_list (embed_list l @ embed_list m)"
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4968
  "[] = return_list []"
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4969
  unfolding return_list_def embed_list_def
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4970
  apply auto
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4971
  apply (induct l, auto)
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4972
  apply (induct m, auto)
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4973
done
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4974
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4975
(*
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4976
lemma transfer_nat_int_fold1: "fold f l x =
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4977
    fold (%x. f (nat x)) (embed_list l) x";
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4978
*)
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4979
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 32960
diff changeset
  4980
37605
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  4981
subsection {* Code generation *}
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  4982
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  4983
subsubsection {* Counterparts for set-related operations *}
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  4984
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  4985
definition member :: "'a list \<Rightarrow> 'a \<Rightarrow> bool" where
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  4986
  [code_post]: "member xs x \<longleftrightarrow> x \<in> set xs"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  4987
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  4988
text {*
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  4989
  Only use @{text member} for generating executable code.  Otherwise use
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  4990
  @{prop "x \<in> set xs"} instead --- it is much easier to reason about.
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  4991
*}
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  4992
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  4993
lemma member_set:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  4994
  "member = set"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39272
diff changeset
  4995
  by (simp add: fun_eq_iff member_def mem_def)
37605
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  4996
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  4997
lemma member_rec [code]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  4998
  "member (x # xs) y \<longleftrightarrow> x = y \<or> member xs y"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  4999
  "member [] y \<longleftrightarrow> False"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5000
  by (auto simp add: member_def)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5001
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5002
lemma in_set_member [code_unfold]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5003
  "x \<in> set xs \<longleftrightarrow> member xs x"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5004
  by (simp add: member_def)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5005
44928
7ef6505bde7f renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents: 44921
diff changeset
  5006
declare INF_def [code_unfold]
7ef6505bde7f renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents: 44921
diff changeset
  5007
declare SUP_def [code_unfold]
37605
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5008
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5009
declare set_map [symmetric, code_unfold]
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5010
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5011
definition list_all :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> bool" where
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5012
  list_all_iff [code_post]: "list_all P xs \<longleftrightarrow> (\<forall>x \<in> set xs. P x)"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5013
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5014
definition list_ex :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> bool" where
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5015
  list_ex_iff [code_post]: "list_ex P xs \<longleftrightarrow> (\<exists>x \<in> set xs. P x)"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5016
40652
7bdfc1d6b143 adding code equations for EX1 on finite types
bulwahn
parents: 40608
diff changeset
  5017
definition list_ex1
7bdfc1d6b143 adding code equations for EX1 on finite types
bulwahn
parents: 40608
diff changeset
  5018
where
7bdfc1d6b143 adding code equations for EX1 on finite types
bulwahn
parents: 40608
diff changeset
  5019
  list_ex1_iff: "list_ex1 P xs \<longleftrightarrow> (\<exists>! x. x \<in> set xs \<and> P x)"
7bdfc1d6b143 adding code equations for EX1 on finite types
bulwahn
parents: 40608
diff changeset
  5020
37605
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5021
text {*
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5022
  Usually you should prefer @{text "\<forall>x\<in>set xs"} and @{text "\<exists>x\<in>set xs"}
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5023
  over @{const list_all} and @{const list_ex} in specifications.
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5024
*}
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5025
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5026
lemma list_all_simps [simp, code]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5027
  "list_all P (x # xs) \<longleftrightarrow> P x \<and> list_all P xs"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5028
  "list_all P [] \<longleftrightarrow> True"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5029
  by (simp_all add: list_all_iff)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5030
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5031
lemma list_ex_simps [simp, code]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5032
  "list_ex P (x # xs) \<longleftrightarrow> P x \<or> list_ex P xs"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5033
  "list_ex P [] \<longleftrightarrow> False"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5034
  by (simp_all add: list_ex_iff)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5035
40652
7bdfc1d6b143 adding code equations for EX1 on finite types
bulwahn
parents: 40608
diff changeset
  5036
lemma list_ex1_simps [simp, code]:
7bdfc1d6b143 adding code equations for EX1 on finite types
bulwahn
parents: 40608
diff changeset
  5037
  "list_ex1 P [] = False"
7bdfc1d6b143 adding code equations for EX1 on finite types
bulwahn
parents: 40608
diff changeset
  5038
  "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)"
7bdfc1d6b143 adding code equations for EX1 on finite types
bulwahn
parents: 40608
diff changeset
  5039
unfolding list_ex1_iff list_all_iff by auto
7bdfc1d6b143 adding code equations for EX1 on finite types
bulwahn
parents: 40608
diff changeset
  5040
37605
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5041
lemma Ball_set_list_all [code_unfold]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5042
  "Ball (set xs) P \<longleftrightarrow> list_all P xs"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5043
  by (simp add: list_all_iff)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5044
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5045
lemma Bex_set_list_ex [code_unfold]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5046
  "Bex (set xs) P \<longleftrightarrow> list_ex P xs"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5047
  by (simp add: list_ex_iff)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5048
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5049
lemma list_all_append [simp]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5050
  "list_all P (xs @ ys) \<longleftrightarrow> list_all P xs \<and> list_all P ys"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5051
  by (auto simp add: list_all_iff)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5052
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5053
lemma list_ex_append [simp]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5054
  "list_ex P (xs @ ys) \<longleftrightarrow> list_ex P xs \<or> list_ex P ys"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5055
  by (auto simp add: list_ex_iff)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5056
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5057
lemma list_all_rev [simp]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5058
  "list_all P (rev xs) \<longleftrightarrow> list_all P xs"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5059
  by (simp add: list_all_iff)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5060
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5061
lemma list_ex_rev [simp]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5062
  "list_ex P (rev xs) \<longleftrightarrow> list_ex P xs"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5063
  by (simp add: list_ex_iff)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5064
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5065
lemma list_all_length:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5066
  "list_all P xs \<longleftrightarrow> (\<forall>n < length xs. P (xs ! n))"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5067
  by (auto simp add: list_all_iff set_conv_nth)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5068
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5069
lemma list_ex_length:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5070
  "list_ex P xs \<longleftrightarrow> (\<exists>n < length xs. P (xs ! n))"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5071
  by (auto simp add: list_ex_iff set_conv_nth)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5072
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5073
lemma list_all_cong [fundef_cong]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5074
  "xs = ys \<Longrightarrow> (\<And>x. x \<in> set ys \<Longrightarrow> f x = g x) \<Longrightarrow> list_all f xs = list_all g ys"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5075
  by (simp add: list_all_iff)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5076
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5077
lemma list_any_cong [fundef_cong]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5078
  "xs = ys \<Longrightarrow> (\<And>x. x \<in> set ys \<Longrightarrow> f x = g x) \<Longrightarrow> list_ex f xs = list_ex g ys"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5079
  by (simp add: list_ex_iff)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5080
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5081
text {* Bounded quantification and summation over nats. *}
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5082
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5083
lemma atMost_upto [code_unfold]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5084
  "{..n} = set [0..<Suc n]"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5085
  by auto
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5086
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5087
lemma atLeast_upt [code_unfold]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5088
  "{..<n} = set [0..<n]"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5089
  by auto
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5090
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5091
lemma greaterThanLessThan_upt [code_unfold]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5092
  "{n<..<m} = set [Suc n..<m]"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5093
  by auto
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5094
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5095
lemmas atLeastLessThan_upt [code_unfold] = set_upt [symmetric]
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5096
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5097
lemma greaterThanAtMost_upt [code_unfold]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5098
  "{n<..m} = set [Suc n..<Suc m]"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5099
  by auto
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5100
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5101
lemma atLeastAtMost_upt [code_unfold]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5102
  "{n..m} = set [n..<Suc m]"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5103
  by auto
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5104
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5105
lemma all_nat_less_eq [code_unfold]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5106
  "(\<forall>m<n\<Colon>nat. P m) \<longleftrightarrow> (\<forall>m \<in> {0..<n}. P m)"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5107
  by auto
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5108
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5109
lemma ex_nat_less_eq [code_unfold]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5110
  "(\<exists>m<n\<Colon>nat. P m) \<longleftrightarrow> (\<exists>m \<in> {0..<n}. P m)"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5111
  by auto
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5112
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5113
lemma all_nat_less [code_unfold]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5114
  "(\<forall>m\<le>n\<Colon>nat. P m) \<longleftrightarrow> (\<forall>m \<in> {0..n}. P m)"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5115
  by auto
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5116
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5117
lemma ex_nat_less [code_unfold]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5118
  "(\<exists>m\<le>n\<Colon>nat. P m) \<longleftrightarrow> (\<exists>m \<in> {0..n}. P m)"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5119
  by auto
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5120
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5121
lemma setsum_set_upt_conv_listsum_nat [code_unfold]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5122
  "setsum f (set [m..<n]) = listsum (map f [m..<n])"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5123
  by (simp add: interv_listsum_conv_setsum_set_nat)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5124
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5125
text {* Summation over ints. *}
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5126
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5127
lemma greaterThanLessThan_upto [code_unfold]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5128
  "{i<..<j::int} = set [i+1..j - 1]"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5129
by auto
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5130
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5131
lemma atLeastLessThan_upto [code_unfold]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5132
  "{i..<j::int} = set [i..j - 1]"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5133
by auto
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5134
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5135
lemma greaterThanAtMost_upto [code_unfold]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5136
  "{i<..j::int} = set [i+1..j]"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5137
by auto
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5138
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5139
lemmas atLeastAtMost_upto [code_unfold] = set_upto [symmetric]
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5140
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5141
lemma setsum_set_upto_conv_listsum_int [code_unfold]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5142
  "setsum f (set [i..j::int]) = listsum (map f [i..j])"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5143
  by (simp add: interv_listsum_conv_setsum_set_int)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5144
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5145
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5146
subsubsection {* Optimizing by rewriting *}
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5147
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5148
definition null :: "'a list \<Rightarrow> bool" where
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5149
  [code_post]: "null xs \<longleftrightarrow> xs = []"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5150
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5151
text {*
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5152
  Efficient emptyness check is implemented by @{const null}.
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5153
*}
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5154
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5155
lemma null_rec [code]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5156
  "null (x # xs) \<longleftrightarrow> False"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5157
  "null [] \<longleftrightarrow> True"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5158
  by (simp_all add: null_def)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5159
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5160
lemma eq_Nil_null [code_unfold]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5161
  "xs = [] \<longleftrightarrow> null xs"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5162
  by (simp add: null_def)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5163
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5164
lemma equal_Nil_null [code_unfold]:
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38715
diff changeset
  5165
  "HOL.equal xs [] \<longleftrightarrow> null xs"
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38715
diff changeset
  5166
  by (simp add: equal eq_Nil_null)
37605
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5167
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5168
definition maps :: "('a \<Rightarrow> 'b list) \<Rightarrow> 'a list \<Rightarrow> 'b list" where
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5169
  [code_post]: "maps f xs = concat (map f xs)"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5170
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5171
definition map_filter :: "('a \<Rightarrow> 'b option) \<Rightarrow> 'a list \<Rightarrow> 'b list" where
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5172
  [code_post]: "map_filter f xs = map (the \<circ> f) (filter (\<lambda>x. f x \<noteq> None) xs)"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5173
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5174
text {*
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5175
  Operations @{const maps} and @{const map_filter} avoid
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5176
  intermediate lists on execution -- do not use for proving.
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5177
*}
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5178
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5179
lemma maps_simps [code]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5180
  "maps f (x # xs) = f x @ maps f xs"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5181
  "maps f [] = []"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5182
  by (simp_all add: maps_def)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5183
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5184
lemma map_filter_simps [code]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5185
  "map_filter f (x # xs) = (case f x of None \<Rightarrow> map_filter f xs | Some y \<Rightarrow> y # map_filter f xs)"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5186
  "map_filter f [] = []"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5187
  by (simp_all add: map_filter_def split: option.split)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5188
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5189
lemma concat_map_maps [code_unfold]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5190
  "concat (map f xs) = maps f xs"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5191
  by (simp add: maps_def)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5192
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5193
lemma map_filter_map_filter [code_unfold]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5194
  "map f (filter P xs) = map_filter (\<lambda>x. if P x then Some (f x) else None) xs"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5195
  by (simp add: map_filter_def)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5196
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5197
text {* Optimized code for @{text"\<forall>i\<in>{a..b::int}"} and @{text"\<forall>n:{a..<b::nat}"}
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5198
and similiarly for @{text"\<exists>"}. *}
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5199
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5200
definition all_interval_nat :: "(nat \<Rightarrow> bool) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5201
  "all_interval_nat P i j \<longleftrightarrow> (\<forall>n \<in> {i..<j}. P n)"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5202
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5203
lemma [code]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5204
  "all_interval_nat P i j \<longleftrightarrow> i \<ge> j \<or> P i \<and> all_interval_nat P (Suc i) j"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5205
proof -
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5206
  have *: "\<And>n. P i \<Longrightarrow> \<forall>n\<in>{Suc i..<j}. P n \<Longrightarrow> i \<le> n \<Longrightarrow> n < j \<Longrightarrow> P n"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5207
  proof -
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5208
    fix n
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5209
    assume "P i" "\<forall>n\<in>{Suc i..<j}. P n" "i \<le> n" "n < j"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5210
    then show "P n" by (cases "n = i") simp_all
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5211
  qed
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5212
  show ?thesis by (auto simp add: all_interval_nat_def intro: *)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5213
qed
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5214
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5215
lemma list_all_iff_all_interval_nat [code_unfold]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5216
  "list_all P [i..<j] \<longleftrightarrow> all_interval_nat P i j"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5217
  by (simp add: list_all_iff all_interval_nat_def)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5218
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5219
lemma list_ex_iff_not_all_inverval_nat [code_unfold]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5220
  "list_ex P [i..<j] \<longleftrightarrow> \<not> (all_interval_nat (Not \<circ> P) i j)"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5221
  by (simp add: list_ex_iff all_interval_nat_def)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5222
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5223
definition all_interval_int :: "(int \<Rightarrow> bool) \<Rightarrow> int \<Rightarrow> int \<Rightarrow> bool" where
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5224
  "all_interval_int P i j \<longleftrightarrow> (\<forall>k \<in> {i..j}. P k)"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5225
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5226
lemma [code]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5227
  "all_interval_int P i j \<longleftrightarrow> i > j \<or> P i \<and> all_interval_int P (i + 1) j"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5228
proof -
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5229
  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"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5230
  proof -
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5231
    fix k
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5232
    assume "P i" "\<forall>k\<in>{i+1..j}. P k" "i \<le> k" "k \<le> j"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5233
    then show "P k" by (cases "k = i") simp_all
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5234
  qed
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5235
  show ?thesis by (auto simp add: all_interval_int_def intro: *)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5236
qed
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5237
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5238
lemma list_all_iff_all_interval_int [code_unfold]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5239
  "list_all P [i..j] \<longleftrightarrow> all_interval_int P i j"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5240
  by (simp add: list_all_iff all_interval_int_def)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5241
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5242
lemma list_ex_iff_not_all_inverval_int [code_unfold]:
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5243
  "list_ex P [i..j] \<longleftrightarrow> \<not> (all_interval_int (Not \<circ> P) i j)"
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5244
  by (simp add: list_ex_iff all_interval_int_def)
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5245
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5246
hide_const (open) member null maps map_filter all_interval_nat all_interval_int
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5247
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5248
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5249
subsubsection {* Pretty lists *}
15064
4f3102b50197 - Moved code generator setup for lists from Main.thy to List.thy
berghofe
parents: 15045
diff changeset
  5250
31055
2cf6efca6c71 proper structures for list and string code generation stuff
haftmann
parents: 31048
diff changeset
  5251
use "Tools/list_code.ML"
2cf6efca6c71 proper structures for list and string code generation stuff
haftmann
parents: 31048
diff changeset
  5252
31048
ac146fc38b51 refined HOL string theories and corresponding ML fragments
haftmann
parents: 31022
diff changeset
  5253
code_type list
ac146fc38b51 refined HOL string theories and corresponding ML fragments
haftmann
parents: 31022
diff changeset
  5254
  (SML "_ list")
ac146fc38b51 refined HOL string theories and corresponding ML fragments
haftmann
parents: 31022
diff changeset
  5255
  (OCaml "_ list")
34886
873c31d9f10d some syntax setup for Scala
haftmann
parents: 34064
diff changeset
  5256
  (Haskell "![(_)]")
873c31d9f10d some syntax setup for Scala
haftmann
parents: 34064
diff changeset
  5257
  (Scala "List[(_)]")
31048
ac146fc38b51 refined HOL string theories and corresponding ML fragments
haftmann
parents: 31022
diff changeset
  5258
ac146fc38b51 refined HOL string theories and corresponding ML fragments
haftmann
parents: 31022
diff changeset
  5259
code_const Nil
ac146fc38b51 refined HOL string theories and corresponding ML fragments
haftmann
parents: 31022
diff changeset
  5260
  (SML "[]")
ac146fc38b51 refined HOL string theories and corresponding ML fragments
haftmann
parents: 31022
diff changeset
  5261
  (OCaml "[]")
ac146fc38b51 refined HOL string theories and corresponding ML fragments
haftmann
parents: 31022
diff changeset
  5262
  (Haskell "[]")
37880
3b9ca8d2c5fb Scala: subtle difference in printing strings vs. complex mixfix syntax
haftmann
parents: 37767
diff changeset
  5263
  (Scala "!Nil")
31048
ac146fc38b51 refined HOL string theories and corresponding ML fragments
haftmann
parents: 31022
diff changeset
  5264
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38715
diff changeset
  5265
code_instance list :: equal
31048
ac146fc38b51 refined HOL string theories and corresponding ML fragments
haftmann
parents: 31022
diff changeset
  5266
  (Haskell -)
ac146fc38b51 refined HOL string theories and corresponding ML fragments
haftmann
parents: 31022
diff changeset
  5267
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38715
diff changeset
  5268
code_const "HOL.equal \<Colon> 'a list \<Rightarrow> 'a list \<Rightarrow> bool"
39272
0b61951d2682 Haskell == is infix, not infixl
haftmann
parents: 39198
diff changeset
  5269
  (Haskell infix 4 "==")
31048
ac146fc38b51 refined HOL string theories and corresponding ML fragments
haftmann
parents: 31022
diff changeset
  5270
ac146fc38b51 refined HOL string theories and corresponding ML fragments
haftmann
parents: 31022
diff changeset
  5271
code_reserved SML
ac146fc38b51 refined HOL string theories and corresponding ML fragments
haftmann
parents: 31022
diff changeset
  5272
  list
ac146fc38b51 refined HOL string theories and corresponding ML fragments
haftmann
parents: 31022
diff changeset
  5273
ac146fc38b51 refined HOL string theories and corresponding ML fragments
haftmann
parents: 31022
diff changeset
  5274
code_reserved OCaml
ac146fc38b51 refined HOL string theories and corresponding ML fragments
haftmann
parents: 31022
diff changeset
  5275
  list
ac146fc38b51 refined HOL string theories and corresponding ML fragments
haftmann
parents: 31022
diff changeset
  5276
45181
c8eb935e2e87 removing old code generator setup for lists
bulwahn
parents: 45115
diff changeset
  5277
setup {* fold (List_Code.add_literal_list) ["SML", "OCaml", "Haskell", "Scala"] *}
15064
4f3102b50197 - Moved code generator setup for lists from Main.thy to List.thy
berghofe
parents: 15045
diff changeset
  5278
21061
580dfc999ef6 added normal post setup; cleaned up "execution" constants
haftmann
parents: 21046
diff changeset
  5279
37424
ed431cc99f17 use various predefined Haskell operations when generating code
haftmann
parents: 37408
diff changeset
  5280
subsubsection {* Use convenient predefined operations *}
ed431cc99f17 use various predefined Haskell operations when generating code
haftmann
parents: 37408
diff changeset
  5281
ed431cc99f17 use various predefined Haskell operations when generating code
haftmann
parents: 37408
diff changeset
  5282
code_const "op @"
ed431cc99f17 use various predefined Haskell operations when generating code
haftmann
parents: 37408
diff changeset
  5283
  (SML infixr 7 "@")
ed431cc99f17 use various predefined Haskell operations when generating code
haftmann
parents: 37408
diff changeset
  5284
  (OCaml infixr 6 "@")
ed431cc99f17 use various predefined Haskell operations when generating code
haftmann
parents: 37408
diff changeset
  5285
  (Haskell infixr 5 "++")
ed431cc99f17 use various predefined Haskell operations when generating code
haftmann
parents: 37408
diff changeset
  5286
  (Scala infixl 7 "++")
ed431cc99f17 use various predefined Haskell operations when generating code
haftmann
parents: 37408
diff changeset
  5287
ed431cc99f17 use various predefined Haskell operations when generating code
haftmann
parents: 37408
diff changeset
  5288
code_const map
ed431cc99f17 use various predefined Haskell operations when generating code
haftmann
parents: 37408
diff changeset
  5289
  (Haskell "map")
ed431cc99f17 use various predefined Haskell operations when generating code
haftmann
parents: 37408
diff changeset
  5290
ed431cc99f17 use various predefined Haskell operations when generating code
haftmann
parents: 37408
diff changeset
  5291
code_const filter
ed431cc99f17 use various predefined Haskell operations when generating code
haftmann
parents: 37408
diff changeset
  5292
  (Haskell "filter")
ed431cc99f17 use various predefined Haskell operations when generating code
haftmann
parents: 37408
diff changeset
  5293
ed431cc99f17 use various predefined Haskell operations when generating code
haftmann
parents: 37408
diff changeset
  5294
code_const concat
ed431cc99f17 use various predefined Haskell operations when generating code
haftmann
parents: 37408
diff changeset
  5295
  (Haskell "concat")
ed431cc99f17 use various predefined Haskell operations when generating code
haftmann
parents: 37408
diff changeset
  5296
37605
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5297
code_const List.maps
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5298
  (Haskell "concatMap")
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5299
37424
ed431cc99f17 use various predefined Haskell operations when generating code
haftmann
parents: 37408
diff changeset
  5300
code_const rev
37451
3058c918e7a3 rev is reverse in Haskell
haftmann
parents: 37424
diff changeset
  5301
  (Haskell "reverse")
37424
ed431cc99f17 use various predefined Haskell operations when generating code
haftmann
parents: 37408
diff changeset
  5302
ed431cc99f17 use various predefined Haskell operations when generating code
haftmann
parents: 37408
diff changeset
  5303
code_const zip
ed431cc99f17 use various predefined Haskell operations when generating code
haftmann
parents: 37408
diff changeset
  5304
  (Haskell "zip")
ed431cc99f17 use various predefined Haskell operations when generating code
haftmann
parents: 37408
diff changeset
  5305
37605
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5306
code_const List.null
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5307
  (Haskell "null")
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5308
37424
ed431cc99f17 use various predefined Haskell operations when generating code
haftmann
parents: 37408
diff changeset
  5309
code_const takeWhile
ed431cc99f17 use various predefined Haskell operations when generating code
haftmann
parents: 37408
diff changeset
  5310
  (Haskell "takeWhile")
ed431cc99f17 use various predefined Haskell operations when generating code
haftmann
parents: 37408
diff changeset
  5311
ed431cc99f17 use various predefined Haskell operations when generating code
haftmann
parents: 37408
diff changeset
  5312
code_const dropWhile
ed431cc99f17 use various predefined Haskell operations when generating code
haftmann
parents: 37408
diff changeset
  5313
  (Haskell "dropWhile")
ed431cc99f17 use various predefined Haskell operations when generating code
haftmann
parents: 37408
diff changeset
  5314
37605
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5315
code_const list_all
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5316
  (Haskell "all")
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5317
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5318
code_const list_ex
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5319
  (Haskell "any")
625bc011768a put section on distinctness before listsum; refined code generation operations; dropped ancient infix mem
haftmann
parents: 37465
diff changeset
  5320
23388
77645da0db85 tuned proofs: avoid implicit prems;
wenzelm
parents: 23279
diff changeset
  5321
end