src/ZF/List.thy
author kleing
Sun, 06 Apr 2003 21:16:50 +0200
changeset 13901 af38553e61ee
parent 13784 b9f6154427a4
child 14046 6616e6c53d48
permissions -rw-r--r--
use 2 processors on sunbroy1
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     1
(*  Title:      ZF/List
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
     2
    ID:         $Id$
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
     4
    Copyright   1994  University of Cambridge
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
     5
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
     6
*)
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
     7
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
     8
header{*Lists in Zermelo-Fraenkel Set Theory*}
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
     9
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
    10
theory List = Datatype + ArithSimp:
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
    11
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
    12
consts
12789
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    13
  list       :: "i=>i"
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
    14
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
    15
datatype
581
465075fd257b removal of needless quotes
lcp
parents: 516
diff changeset
    16
  "list(A)" = Nil | Cons ("a:A", "l: list(A)")
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
    17
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
    18
2539
ddd22ceee8cc turned some consts into syntax;
wenzelm
parents: 1926
diff changeset
    19
syntax
ddd22ceee8cc turned some consts into syntax;
wenzelm
parents: 1926
diff changeset
    20
 "[]"        :: i                                       ("[]")
12789
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    21
 "@List"     :: "is => i"                                 ("[(_)]")
2539
ddd22ceee8cc turned some consts into syntax;
wenzelm
parents: 1926
diff changeset
    22
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
    23
translations
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
    24
  "[x, xs]"     == "Cons(x, [xs])"
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
    25
  "[x]"         == "Cons(x, [])"
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
    26
  "[]"          == "Nil"
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
    27
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
    28
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    29
consts
12789
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    30
  length :: "i=>i"
13396
11219ca224ab A couple of new theorems for Constructible
paulson
parents: 13387
diff changeset
    31
  hd     :: "i=>i"
11219ca224ab A couple of new theorems for Constructible
paulson
parents: 13387
diff changeset
    32
  tl     :: "i=>i"
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    33
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    34
primrec
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    35
  "length([]) = 0"
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    36
  "length(Cons(a,l)) = succ(length(l))"
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
    37
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    38
primrec
13396
11219ca224ab A couple of new theorems for Constructible
paulson
parents: 13387
diff changeset
    39
  "hd([]) = 0"
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    40
  "hd(Cons(a,l)) = a"
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    41
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    42
primrec
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    43
  "tl([]) = []"
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    44
  "tl(Cons(a,l)) = l"
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    45
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6053
diff changeset
    46
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6053
diff changeset
    47
consts
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
    48
  map         :: "[i=>i, i] => i"
12789
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    49
  set_of_list :: "i=>i"
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
    50
  app         :: "[i,i]=>i"                        (infixr "@" 60)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
    51
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
    52
(*map is a binding operator -- it applies to meta-level functions, not
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
    53
object-level functions.  This simplifies the final form of term_rec_conv,
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
    54
although complicating its derivation.*)
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    55
primrec
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    56
  "map(f,[]) = []"
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    57
  "map(f,Cons(a,l)) = Cons(f(a), map(f,l))"
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
    58
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    59
primrec
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    60
  "set_of_list([]) = 0"
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    61
  "set_of_list(Cons(a,l)) = cons(a, set_of_list(l))"
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
    62
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    63
primrec
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
    64
  app_Nil:  "[] @ ys = ys"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
    65
  app_Cons: "(Cons(a,l)) @ ys = Cons(a, l @ ys)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
    66
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6053
diff changeset
    67
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6053
diff changeset
    68
consts
12789
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    69
  rev :: "i=>i"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    70
  flat       :: "i=>i"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    71
  list_add   :: "i=>i"
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6053
diff changeset
    72
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    73
primrec
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    74
  "rev([]) = []"
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    75
  "rev(Cons(a,l)) = rev(l) @ [a]"
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    76
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    77
primrec
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    78
  "flat([]) = []"
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    79
  "flat(Cons(l,ls)) = l @ flat(ls)"
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
    80
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    81
primrec
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    82
  "list_add([]) = 0"
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    83
  "list_add(Cons(a,l)) = a #+ list_add(l)"
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
    84
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6053
diff changeset
    85
consts
12789
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    86
  drop       :: "[i,i]=>i"
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6053
diff changeset
    87
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6053
diff changeset
    88
primrec
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
    89
  drop_0:    "drop(0,l) = l"
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
    90
  drop_succ: "drop(succ(i), l) = tl (drop(i,l))"
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
    91
12789
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    92
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    93
(*** Thanks to Sidi Ehmety for the following ***)
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    94
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    95
constdefs
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    96
(* Function `take' returns the first n elements of a list *)
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    97
  take     :: "[i,i]=>i"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    98
  "take(n, as) == list_rec(lam n:nat. [],
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    99
		%a l r. lam n:nat. nat_case([], %m. Cons(a, r`m), n), as)`n"
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   100
12789
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   101
  nth :: "[i, i]=>i"
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   102
  --{*returns the (n+1)th element of a list, or 0 if the
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   103
   list is too short.*}
12789
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   104
  "nth(n, as) == list_rec(lam n:nat. 0,
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   105
 		          %a l r. lam n:nat. nat_case(a, %m. r`m, n), as) ` n"
12789
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   106
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   107
  list_update :: "[i, i, i]=>i"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   108
  "list_update(xs, i, v) == list_rec(lam n:nat. Nil,
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   109
      %u us vs. lam n:nat. nat_case(Cons(v, us), %m. Cons(u, vs`m), n), xs)`i"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   110
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   111
consts
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   112
  filter :: "[i=>o, i] => i"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   113
  upt :: "[i, i] =>i"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   114
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   115
primrec
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   116
  "filter(P, Nil) = Nil"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   117
  "filter(P, Cons(x, xs)) =
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   118
     (if P(x) then Cons(x, filter(P, xs)) else filter(P, xs))"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   119
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   120
primrec
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   121
  "upt(i, 0) = Nil"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   122
  "upt(i, succ(j)) = (if i le j then upt(i, j)@[j] else Nil)"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   123
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   124
constdefs
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   125
  min :: "[i,i] =>i"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   126
    "min(x, y) == (if x le y then x else y)"
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   127
12789
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   128
  max :: "[i, i] =>i"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   129
    "max(x, y) == (if x le y then y else x)"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   130
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   131
(*** Aspects of the datatype definition ***)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   132
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   133
declare list.intros [simp,TC]
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   134
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   135
(*An elimination rule, for type-checking*)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   136
inductive_cases ConsE: "Cons(a,l) : list(A)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   137
13509
6f168374652a new lemmas
paulson
parents: 13396
diff changeset
   138
lemma Cons_type_iff [simp]: "Cons(a,l) \<in> list(A) <-> a \<in> A & l \<in> list(A)"
6f168374652a new lemmas
paulson
parents: 13396
diff changeset
   139
by (blast elim: ConsE) 
6f168374652a new lemmas
paulson
parents: 13396
diff changeset
   140
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   141
(*Proving freeness results*)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   142
lemma Cons_iff: "Cons(a,l)=Cons(a',l') <-> a=a' & l=l'"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   143
by auto
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   144
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   145
lemma Nil_Cons_iff: "~ Nil=Cons(a,l)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   146
by auto
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   147
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   148
lemma list_unfold: "list(A) = {0} + (A * list(A))"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   149
by (blast intro!: list.intros [unfolded list.con_defs]
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   150
          elim: list.cases [unfolded list.con_defs])
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   151
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   152
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   153
(**  Lemmas to justify using "list" in other recursive type definitions **)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   154
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   155
lemma list_mono: "A<=B ==> list(A) <= list(B)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   156
apply (unfold list.defs )
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   157
apply (rule lfp_mono)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   158
apply (simp_all add: list.bnd_mono)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   159
apply (assumption | rule univ_mono basic_monos)+
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   160
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   161
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   162
(*There is a similar proof by list induction.*)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   163
lemma list_univ: "list(univ(A)) <= univ(A)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   164
apply (unfold list.defs list.con_defs)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   165
apply (rule lfp_lowerbound)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   166
apply (rule_tac [2] A_subset_univ [THEN univ_mono])
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   167
apply (blast intro!: zero_in_univ Inl_in_univ Inr_in_univ Pair_in_univ)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   168
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   169
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   170
(*These two theorems justify datatypes involving list(nat), list(A), ...*)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   171
lemmas list_subset_univ = subset_trans [OF list_mono list_univ]
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   172
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   173
lemma list_into_univ: "[| l: list(A);  A <= univ(B) |] ==> l: univ(B)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   174
by (blast intro: list_subset_univ [THEN subsetD])
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   175
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   176
lemma list_case_type:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   177
    "[| l: list(A);
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   178
        c: C(Nil);
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   179
        !!x y. [| x: A;  y: list(A) |] ==> h(x,y): C(Cons(x,y))
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   180
     |] ==> list_case(c,h,l) : C(l)"
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   181
by (erule list.induct, auto)
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   182
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   183
lemma list_0_triv: "list(0) = {Nil}"
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   184
apply (rule equalityI, auto) 
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   185
apply (induct_tac x, auto) 
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   186
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   187
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   188
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   189
(*** List functions ***)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   190
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   191
lemma tl_type: "l: list(A) ==> tl(l) : list(A)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   192
apply (induct_tac "l")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   193
apply (simp_all (no_asm_simp) add: list.intros)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   194
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   195
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   196
(** drop **)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   197
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   198
lemma drop_Nil [simp]: "i:nat ==> drop(i, Nil) = Nil"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   199
apply (induct_tac "i")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   200
apply (simp_all (no_asm_simp))
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   201
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   202
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   203
lemma drop_succ_Cons [simp]: "i:nat ==> drop(succ(i), Cons(a,l)) = drop(i,l)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   204
apply (rule sym)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   205
apply (induct_tac "i")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   206
apply (simp (no_asm))
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   207
apply (simp (no_asm_simp))
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   208
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   209
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   210
lemma drop_type [simp,TC]: "[| i:nat; l: list(A) |] ==> drop(i,l) : list(A)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   211
apply (induct_tac "i")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   212
apply (simp_all (no_asm_simp) add: tl_type)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   213
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   214
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   215
declare drop_succ [simp del]
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   216
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   217
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   218
(** Type checking -- proved by induction, as usual **)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   219
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   220
lemma list_rec_type [TC]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   221
    "[| l: list(A);
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   222
        c: C(Nil);
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   223
        !!x y r. [| x:A;  y: list(A);  r: C(y) |] ==> h(x,y,r): C(Cons(x,y))
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   224
     |] ==> list_rec(c,h,l) : C(l)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   225
by (induct_tac "l", auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   226
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   227
(** map **)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   228
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   229
lemma map_type [TC]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   230
    "[| l: list(A);  !!x. x: A ==> h(x): B |] ==> map(h,l) : list(B)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   231
apply (simp add: map_list_def)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   232
apply (typecheck add: list.intros list_rec_type, blast)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   233
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   234
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   235
lemma map_type2 [TC]: "l: list(A) ==> map(h,l) : list({h(u). u:A})"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   236
apply (erule map_type)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   237
apply (erule RepFunI)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   238
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   239
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   240
(** length **)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   241
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   242
lemma length_type [TC]: "l: list(A) ==> length(l) : nat"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   243
by (simp add: length_list_def)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   244
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   245
lemma lt_length_in_nat:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   246
   "[|x < length(xs); xs \<in> list(A)|] ==> x \<in> nat"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   247
by (frule lt_nat_in_nat, typecheck) 
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   248
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   249
(** app **)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   250
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   251
lemma app_type [TC]: "[| xs: list(A);  ys: list(A) |] ==> xs@ys : list(A)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   252
by (simp add: app_list_def)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   253
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   254
(** rev **)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   255
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   256
lemma rev_type [TC]: "xs: list(A) ==> rev(xs) : list(A)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   257
by (simp add: rev_list_def)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   258
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   259
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   260
(** flat **)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   261
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   262
lemma flat_type [TC]: "ls: list(list(A)) ==> flat(ls) : list(A)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   263
by (simp add: flat_list_def)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   264
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   265
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   266
(** set_of_list **)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   267
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   268
lemma set_of_list_type [TC]: "l: list(A) ==> set_of_list(l) : Pow(A)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   269
apply (unfold set_of_list_list_def)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   270
apply (erule list_rec_type, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   271
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   272
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   273
lemma set_of_list_append:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   274
     "xs: list(A) ==> set_of_list (xs@ys) = set_of_list(xs) Un set_of_list(ys)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   275
apply (erule list.induct)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   276
apply (simp_all (no_asm_simp) add: Un_cons)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   277
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   278
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   279
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   280
(** list_add **)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   281
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   282
lemma list_add_type [TC]: "xs: list(nat) ==> list_add(xs) : nat"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   283
by (simp add: list_add_list_def)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   284
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   285
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   286
(*** theorems about map ***)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   287
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   288
lemma map_ident [simp]: "l: list(A) ==> map(%u. u, l) = l"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   289
apply (induct_tac "l")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   290
apply (simp_all (no_asm_simp))
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   291
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   292
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   293
lemma map_compose: "l: list(A) ==> map(h, map(j,l)) = map(%u. h(j(u)), l)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   294
apply (induct_tac "l")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   295
apply (simp_all (no_asm_simp))
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   296
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   297
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   298
lemma map_app_distrib: "xs: list(A) ==> map(h, xs@ys) = map(h,xs) @ map(h,ys)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   299
apply (induct_tac "xs")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   300
apply (simp_all (no_asm_simp))
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   301
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   302
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   303
lemma map_flat: "ls: list(list(A)) ==> map(h, flat(ls)) = flat(map(map(h),ls))"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   304
apply (induct_tac "ls")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   305
apply (simp_all (no_asm_simp) add: map_app_distrib)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   306
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   307
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   308
lemma list_rec_map:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   309
     "l: list(A) ==>
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   310
      list_rec(c, d, map(h,l)) =
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   311
      list_rec(c, %x xs r. d(h(x), map(h,xs), r), l)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   312
apply (induct_tac "l")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   313
apply (simp_all (no_asm_simp))
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   314
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   315
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   316
(** theorems about list(Collect(A,P)) -- used in Induct/Term.thy **)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   317
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   318
(* c : list(Collect(B,P)) ==> c : list(B) *)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   319
lemmas list_CollectD = Collect_subset [THEN list_mono, THEN subsetD, standard]
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   320
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   321
lemma map_list_Collect: "l: list({x:A. h(x)=j(x)}) ==> map(h,l) = map(j,l)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   322
apply (induct_tac "l")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   323
apply (simp_all (no_asm_simp))
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   324
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   325
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   326
(*** theorems about length ***)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   328
lemma length_map [simp]: "xs: list(A) ==> length(map(h,xs)) = length(xs)"
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   329
by (induct_tac "xs", simp_all)
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   330
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   331
lemma length_app [simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   332
     "[| xs: list(A); ys: list(A) |]
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   333
      ==> length(xs@ys) = length(xs) #+ length(ys)"
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   334
by (induct_tac "xs", simp_all)
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   335
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   336
lemma length_rev [simp]: "xs: list(A) ==> length(rev(xs)) = length(xs)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   337
apply (induct_tac "xs")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   338
apply (simp_all (no_asm_simp) add: length_app)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   339
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   340
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   341
lemma length_flat:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   342
     "ls: list(list(A)) ==> length(flat(ls)) = list_add(map(length,ls))"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   343
apply (induct_tac "ls")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   344
apply (simp_all (no_asm_simp) add: length_app)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   345
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   346
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   347
(** Length and drop **)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   348
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   349
(*Lemma for the inductive step of drop_length*)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   350
lemma drop_length_Cons [rule_format]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   351
     "xs: list(A) ==>
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   352
           \<forall>x.  EX z zs. drop(length(xs), Cons(x,xs)) = Cons(z,zs)"
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   353
by (erule list.induct, simp_all)
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   354
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   355
lemma drop_length [rule_format]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   356
     "l: list(A) ==> \<forall>i \<in> length(l). (EX z zs. drop(i,l) = Cons(z,zs))"
13784
b9f6154427a4 tidying (by script)
paulson
parents: 13615
diff changeset
   357
apply (erule list.induct, simp_all, safe)
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   358
apply (erule drop_length_Cons)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   359
apply (rule natE)
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   360
apply (erule Ord_trans [OF asm_rl length_type Ord_nat], assumption, simp_all)
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   361
apply (blast intro: succ_in_naturalD length_type)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   362
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   363
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   364
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   365
(*** theorems about app ***)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   366
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   367
lemma app_right_Nil [simp]: "xs: list(A) ==> xs@Nil=xs"
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   368
by (erule list.induct, simp_all)
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   369
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   370
lemma app_assoc: "xs: list(A) ==> (xs@ys)@zs = xs@(ys@zs)"
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   371
by (induct_tac "xs", simp_all)
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   372
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   373
lemma flat_app_distrib: "ls: list(list(A)) ==> flat(ls@ms) = flat(ls)@flat(ms)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   374
apply (induct_tac "ls")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   375
apply (simp_all (no_asm_simp) add: app_assoc)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   376
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   377
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   378
(*** theorems about rev ***)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   379
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   380
lemma rev_map_distrib: "l: list(A) ==> rev(map(h,l)) = map(h,rev(l))"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   381
apply (induct_tac "l")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   382
apply (simp_all (no_asm_simp) add: map_app_distrib)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   383
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   384
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   385
(*Simplifier needs the premises as assumptions because rewriting will not
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   386
  instantiate the variable ?A in the rules' typing conditions; note that
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   387
  rev_type does not instantiate ?A.  Only the premises do.
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   388
*)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   389
lemma rev_app_distrib:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   390
     "[| xs: list(A);  ys: list(A) |] ==> rev(xs@ys) = rev(ys)@rev(xs)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   391
apply (erule list.induct)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   392
apply (simp_all add: app_assoc)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   393
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   394
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   395
lemma rev_rev_ident [simp]: "l: list(A) ==> rev(rev(l))=l"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   396
apply (induct_tac "l")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   397
apply (simp_all (no_asm_simp) add: rev_app_distrib)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   398
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   399
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   400
lemma rev_flat: "ls: list(list(A)) ==> rev(flat(ls)) = flat(map(rev,rev(ls)))"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   401
apply (induct_tac "ls")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   402
apply (simp_all add: map_app_distrib flat_app_distrib rev_app_distrib)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   403
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   404
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   405
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   406
(*** theorems about list_add ***)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   407
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   408
lemma list_add_app:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   409
     "[| xs: list(nat);  ys: list(nat) |]
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   410
      ==> list_add(xs@ys) = list_add(ys) #+ list_add(xs)"
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   411
apply (induct_tac "xs", simp_all)
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   412
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   413
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   414
lemma list_add_rev: "l: list(nat) ==> list_add(rev(l)) = list_add(l)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   415
apply (induct_tac "l")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   416
apply (simp_all (no_asm_simp) add: list_add_app)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   417
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   418
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   419
lemma list_add_flat:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   420
     "ls: list(list(nat)) ==> list_add(flat(ls)) = list_add(map(list_add,ls))"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   421
apply (induct_tac "ls")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   422
apply (simp_all (no_asm_simp) add: list_add_app)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   423
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   424
13509
6f168374652a new lemmas
paulson
parents: 13396
diff changeset
   425
(** New induction rules **)
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   426
13524
604d0f3622d6 *** empty log message ***
wenzelm
parents: 13509
diff changeset
   427
lemma list_append_induct [case_names Nil snoc, consumes 1]:
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   428
    "[| l: list(A);
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   429
        P(Nil);
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   430
        !!x y. [| x: A;  y: list(A);  P(y) |] ==> P(y @ [x])
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   431
     |] ==> P(l)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   432
apply (subgoal_tac "P(rev(rev(l)))", simp)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   433
apply (erule rev_type [THEN list.induct], simp_all)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   434
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   435
13509
6f168374652a new lemmas
paulson
parents: 13396
diff changeset
   436
lemma list_complete_induct_lemma [rule_format]:
6f168374652a new lemmas
paulson
parents: 13396
diff changeset
   437
 assumes ih: 
6f168374652a new lemmas
paulson
parents: 13396
diff changeset
   438
    "\<And>l. [| l \<in> list(A); 
6f168374652a new lemmas
paulson
parents: 13396
diff changeset
   439
             \<forall>l' \<in> list(A). length(l') < length(l) --> P(l')|] 
6f168374652a new lemmas
paulson
parents: 13396
diff changeset
   440
          ==> P(l)"
6f168374652a new lemmas
paulson
parents: 13396
diff changeset
   441
  shows "n \<in> nat ==> \<forall>l \<in> list(A). length(l) < n --> P(l)"
6f168374652a new lemmas
paulson
parents: 13396
diff changeset
   442
apply (induct_tac n, simp)
6f168374652a new lemmas
paulson
parents: 13396
diff changeset
   443
apply (blast intro: ih elim!: leE) 
6f168374652a new lemmas
paulson
parents: 13396
diff changeset
   444
done
6f168374652a new lemmas
paulson
parents: 13396
diff changeset
   445
6f168374652a new lemmas
paulson
parents: 13396
diff changeset
   446
theorem list_complete_induct:
6f168374652a new lemmas
paulson
parents: 13396
diff changeset
   447
      "[| l \<in> list(A); 
6f168374652a new lemmas
paulson
parents: 13396
diff changeset
   448
          \<And>l. [| l \<in> list(A); 
6f168374652a new lemmas
paulson
parents: 13396
diff changeset
   449
                  \<forall>l' \<in> list(A). length(l') < length(l) --> P(l')|] 
6f168374652a new lemmas
paulson
parents: 13396
diff changeset
   450
               ==> P(l)
6f168374652a new lemmas
paulson
parents: 13396
diff changeset
   451
       |] ==> P(l)"
6f168374652a new lemmas
paulson
parents: 13396
diff changeset
   452
apply (rule list_complete_induct_lemma [of A]) 
6f168374652a new lemmas
paulson
parents: 13396
diff changeset
   453
   prefer 4 apply (rule le_refl, simp) 
6f168374652a new lemmas
paulson
parents: 13396
diff changeset
   454
  apply blast 
6f168374652a new lemmas
paulson
parents: 13396
diff changeset
   455
 apply simp 
6f168374652a new lemmas
paulson
parents: 13396
diff changeset
   456
apply assumption
6f168374652a new lemmas
paulson
parents: 13396
diff changeset
   457
done
6f168374652a new lemmas
paulson
parents: 13396
diff changeset
   458
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   459
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   460
(*** Thanks to Sidi Ehmety for these results about min, take, etc. ***)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   461
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   462
(** min FIXME: replace by Int! **)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   463
(* Min theorems are also true for i, j ordinals *)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   464
lemma min_sym: "[| i:nat; j:nat |] ==> min(i,j)=min(j,i)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   465
apply (unfold min_def)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   466
apply (auto dest!: not_lt_imp_le dest: lt_not_sym intro: le_anti_sym)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   467
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   468
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   469
lemma min_type [simp,TC]: "[| i:nat; j:nat |] ==> min(i,j):nat"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   470
by (unfold min_def, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   471
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   472
lemma min_0 [simp]: "i:nat ==> min(0,i) = 0"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   473
apply (unfold min_def)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   474
apply (auto dest: not_lt_imp_le)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   475
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   476
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   477
lemma min_02 [simp]: "i:nat ==> min(i, 0) = 0"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   478
apply (unfold min_def)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   479
apply (auto dest: not_lt_imp_le)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   480
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   481
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   482
lemma lt_min_iff: "[| i:nat; j:nat; k:nat |] ==> i<min(j,k) <-> i<j & i<k"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   483
apply (unfold min_def)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   484
apply (auto dest!: not_lt_imp_le intro: lt_trans2 lt_trans)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   485
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   486
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   487
lemma min_succ_succ [simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   488
     "[| i:nat; j:nat |] ==>  min(succ(i), succ(j))= succ(min(i, j))"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   489
apply (unfold min_def, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   490
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   491
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   492
(*** more theorems about lists ***)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   493
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   494
(** filter **)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   495
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   496
lemma filter_append [simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   497
     "xs:list(A) ==> filter(P, xs@ys) = filter(P, xs) @ filter(P, ys)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   498
by (induct_tac "xs", auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   499
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   500
lemma filter_type [simp,TC]: "xs:list(A) ==> filter(P, xs):list(A)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   501
by (induct_tac "xs", auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   502
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   503
lemma length_filter: "xs:list(A) ==> length(filter(P, xs)) le length(xs)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   504
apply (induct_tac "xs", auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   505
apply (rule_tac j = "length (l) " in le_trans)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   506
apply (auto simp add: le_iff)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   507
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   508
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   509
lemma filter_is_subset: "xs:list(A) ==> set_of_list(filter(P,xs)) <= set_of_list(xs)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   510
by (induct_tac "xs", auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   511
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   512
lemma filter_False [simp]: "xs:list(A) ==> filter(%p. False, xs) = Nil"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   513
by (induct_tac "xs", auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   514
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   515
lemma filter_True [simp]: "xs:list(A) ==> filter(%p. True, xs) = xs"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   516
by (induct_tac "xs", auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   517
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   518
(** length **)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   519
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   520
lemma length_is_0_iff [simp]: "xs:list(A) ==> length(xs)=0 <-> xs=Nil"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   521
by (erule list.induct, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   522
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   523
lemma length_is_0_iff2 [simp]: "xs:list(A) ==> 0 = length(xs) <-> xs=Nil"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   524
by (erule list.induct, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   525
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   526
lemma length_tl [simp]: "xs:list(A) ==> length(tl(xs)) = length(xs) #- 1"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   527
by (erule list.induct, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   528
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   529
lemma length_greater_0_iff: "xs:list(A) ==> 0<length(xs) <-> xs ~= Nil"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   530
by (erule list.induct, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   531
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   532
lemma length_succ_iff: "xs:list(A) ==> length(xs)=succ(n) <-> (EX y ys. xs=Cons(y, ys) & length(ys)=n)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   533
by (erule list.induct, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   534
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   535
(** more theorems about append **)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   536
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   537
lemma append_is_Nil_iff [simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   538
     "xs:list(A) ==> (xs@ys = Nil) <-> (xs=Nil & ys = Nil)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   539
by (erule list.induct, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   540
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   541
lemma append_is_Nil_iff2 [simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   542
     "xs:list(A) ==> (Nil = xs@ys) <-> (xs=Nil & ys = Nil)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   543
by (erule list.induct, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   544
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   545
lemma append_left_is_self_iff [simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   546
     "xs:list(A) ==> (xs@ys = xs) <-> (ys = Nil)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   547
by (erule list.induct, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   548
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   549
lemma append_left_is_self_iff2 [simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   550
     "xs:list(A) ==> (xs = xs@ys) <-> (ys = Nil)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   551
by (erule list.induct, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   552
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   553
(*TOO SLOW as a default simprule!*)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   554
lemma append_left_is_Nil_iff [rule_format]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   555
     "[| xs:list(A); ys:list(A); zs:list(A) |] ==>
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   556
   length(ys)=length(zs) --> (xs@ys=zs <-> (xs=Nil & ys=zs))"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   557
apply (erule list.induct)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   558
apply (auto simp add: length_app)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   559
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   560
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   561
(*TOO SLOW as a default simprule!*)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   562
lemma append_left_is_Nil_iff2 [rule_format]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   563
     "[| xs:list(A); ys:list(A); zs:list(A) |] ==>
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   564
   length(ys)=length(zs) --> (zs=ys@xs <-> (xs=Nil & ys=zs))"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   565
apply (erule list.induct)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   566
apply (auto simp add: length_app)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   567
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   568
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   569
lemma append_eq_append_iff [rule_format,simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   570
     "xs:list(A) ==> \<forall>ys \<in> list(A).
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   571
      length(xs)=length(ys) --> (xs@us = ys@vs) <-> (xs=ys & us=vs)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   572
apply (erule list.induct)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   573
apply (simp (no_asm_simp))
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   574
apply clarify
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   575
apply (erule_tac a = ys in list.cases, auto)
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   576
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   577
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   578
lemma append_eq_append [rule_format]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   579
  "xs:list(A) ==>
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   580
   \<forall>ys \<in> list(A). \<forall>us \<in> list(A). \<forall>vs \<in> list(A).
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   581
   length(us) = length(vs) --> (xs@us = ys@vs) --> (xs=ys & us=vs)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   582
apply (induct_tac "xs")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   583
apply (force simp add: length_app, clarify)
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   584
apply (erule_tac a = ys in list.cases, simp)
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   585
apply (subgoal_tac "Cons (a, l) @ us =vs")
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   586
 apply (drule rev_iffD1 [OF _ append_left_is_Nil_iff], simp_all, blast)
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   587
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   588
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   589
lemma append_eq_append_iff2 [simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   590
 "[| xs:list(A); ys:list(A); us:list(A); vs:list(A); length(us)=length(vs) |]
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   591
  ==>  xs@us = ys@vs <-> (xs=ys & us=vs)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   592
apply (rule iffI)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   593
apply (rule append_eq_append, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   594
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   595
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   596
lemma append_self_iff [simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   597
     "[| xs:list(A); ys:list(A); zs:list(A) |] ==> xs@ys=xs@zs <-> ys=zs"
13509
6f168374652a new lemmas
paulson
parents: 13396
diff changeset
   598
by simp
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   599
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   600
lemma append_self_iff2 [simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   601
     "[| xs:list(A); ys:list(A); zs:list(A) |] ==> ys@xs=zs@xs <-> ys=zs"
13509
6f168374652a new lemmas
paulson
parents: 13396
diff changeset
   602
by simp
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   603
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   604
(* Can also be proved from append_eq_append_iff2,
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   605
but the proof requires two more hypotheses: x:A and y:A *)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   606
lemma append1_eq_iff [rule_format,simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   607
     "xs:list(A) ==> \<forall>ys \<in> list(A). xs@[x] = ys@[y] <-> (xs = ys & x=y)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   608
apply (erule list.induct)  
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   609
 apply clarify 
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   610
 apply (erule list.cases)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   611
 apply simp_all
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   612
txt{*Inductive step*}  
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   613
apply clarify 
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13327
diff changeset
   614
apply (erule_tac a=ys in list.cases, simp_all)  
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   615
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   616
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   617
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   618
lemma append_right_is_self_iff [simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   619
     "[| xs:list(A); ys:list(A) |] ==> (xs@ys = ys) <-> (xs=Nil)"
13509
6f168374652a new lemmas
paulson
parents: 13396
diff changeset
   620
by (simp (no_asm_simp) add: append_left_is_Nil_iff)
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   621
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   622
lemma append_right_is_self_iff2 [simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   623
     "[| xs:list(A); ys:list(A) |] ==> (ys = xs@ys) <-> (xs=Nil)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   624
apply (rule iffI)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   625
apply (drule sym, auto) 
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   626
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   627
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   628
lemma hd_append [rule_format,simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   629
     "xs:list(A) ==> xs ~= Nil --> hd(xs @ ys) = hd(xs)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   630
by (induct_tac "xs", auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   631
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   632
lemma tl_append [rule_format,simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   633
     "xs:list(A) ==> xs~=Nil --> tl(xs @ ys) = tl(xs)@ys"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   634
by (induct_tac "xs", auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   635
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   636
(** rev **)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   637
lemma rev_is_Nil_iff [simp]: "xs:list(A) ==> (rev(xs) = Nil <-> xs = Nil)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   638
by (erule list.induct, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   639
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   640
lemma Nil_is_rev_iff [simp]: "xs:list(A) ==> (Nil = rev(xs) <-> xs = Nil)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   641
by (erule list.induct, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   642
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   643
lemma rev_is_rev_iff [rule_format,simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   644
     "xs:list(A) ==> \<forall>ys \<in> list(A). rev(xs)=rev(ys) <-> xs=ys"
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   645
apply (erule list.induct, force, clarify)
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   646
apply (erule_tac a = ys in list.cases, auto)
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   647
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   648
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   649
lemma rev_list_elim [rule_format]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   650
     "xs:list(A) ==>
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   651
      (xs=Nil --> P) --> (\<forall>ys \<in> list(A). \<forall>y \<in> A. xs =ys@[y] -->P)-->P"
13509
6f168374652a new lemmas
paulson
parents: 13396
diff changeset
   652
by (erule list_append_induct, auto)
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   653
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   654
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   655
(** more theorems about drop **)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   656
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   657
lemma length_drop [rule_format,simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   658
     "n:nat ==> \<forall>xs \<in> list(A). length(drop(n, xs)) = length(xs) #- n"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   659
apply (erule nat_induct)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   660
apply (auto elim: list.cases)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   661
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   662
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   663
lemma drop_all [rule_format,simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   664
     "n:nat ==> \<forall>xs \<in> list(A). length(xs) le n --> drop(n, xs)=Nil"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   665
apply (erule nat_induct)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   666
apply (auto elim: list.cases)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   667
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   668
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   669
lemma drop_append [rule_format]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   670
     "n:nat ==> 
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   671
      \<forall>xs \<in> list(A). drop(n, xs@ys) = drop(n,xs) @ drop(n #- length(xs), ys)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   672
apply (induct_tac "n")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   673
apply (auto elim: list.cases)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   674
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   675
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   676
lemma drop_drop:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   677
    "m:nat ==> \<forall>xs \<in> list(A). \<forall>n \<in> nat. drop(n, drop(m, xs))=drop(n #+ m, xs)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   678
apply (induct_tac "m")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   679
apply (auto elim: list.cases)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   680
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   681
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   682
(** take **)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   683
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   684
lemma take_0 [simp]: "xs:list(A) ==> take(0, xs) =  Nil"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   685
apply (unfold take_def)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   686
apply (erule list.induct, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   687
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   688
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   689
lemma take_succ_Cons [simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   690
    "n:nat ==> take(succ(n), Cons(a, xs)) = Cons(a, take(n, xs))"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   691
by (simp add: take_def)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   692
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   693
(* Needed for proving take_all *)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   694
lemma take_Nil [simp]: "n:nat ==> take(n, Nil) = Nil"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   695
by (unfold take_def, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   696
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   697
lemma take_all [rule_format,simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   698
     "n:nat ==> \<forall>xs \<in> list(A). length(xs) le n  --> take(n, xs) = xs"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   699
apply (erule nat_induct)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   700
apply (auto elim: list.cases) 
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   701
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   702
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   703
lemma take_type [rule_format,simp,TC]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   704
     "xs:list(A) ==> \<forall>n \<in> nat. take(n, xs):list(A)"
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   705
apply (erule list.induct, simp, clarify) 
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   706
apply (erule natE, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   707
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   708
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   709
lemma take_append [rule_format,simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   710
 "xs:list(A) ==>
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   711
  \<forall>ys \<in> list(A). \<forall>n \<in> nat. take(n, xs @ ys) =
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   712
                            take(n, xs) @ take(n #- length(xs), ys)"
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   713
apply (erule list.induct, simp, clarify) 
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   714
apply (erule natE, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   715
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   716
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   717
lemma take_take [rule_format]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   718
   "m : nat ==>
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   719
    \<forall>xs \<in> list(A). \<forall>n \<in> nat. take(n, take(m,xs))= take(min(n, m), xs)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   720
apply (induct_tac "m", auto)
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   721
apply (erule_tac a = xs in list.cases)
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   722
apply (auto simp add: take_Nil)
13615
449a70d88b38 Numerous cosmetic changes, prompted by the new simplifier
paulson
parents: 13611
diff changeset
   723
apply (erule_tac n=n in natE)
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   724
apply (auto intro: take_0 take_type)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   725
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   726
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   727
(** nth **)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   728
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   729
lemma nth_0 [simp]: "nth(0, Cons(a, l)) = a"
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   730
by (simp add: nth_def) 
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   731
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   732
lemma nth_Cons [simp]: "n:nat ==> nth(succ(n), Cons(a,l)) = nth(n,l)"
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   733
by (simp add: nth_def) 
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   734
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   735
lemma nth_empty [simp]: "nth(n, Nil) = 0"
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   736
by (simp add: nth_def) 
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   737
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   738
lemma nth_type [rule_format,simp,TC]:
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   739
     "xs:list(A) ==> \<forall>n \<in> nat. n < length(xs) --> nth(n,xs) : A"
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   740
apply (erule list.induct, simp, clarify) 
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   741
apply (erule natE, auto)
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   742
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   743
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   744
lemma nth_eq_0 [rule_format]:
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   745
     "xs:list(A) ==> \<forall>n \<in> nat. length(xs) le n --> nth(n,xs) = 0"
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   746
apply (erule list.induct, simp, clarify) 
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   747
apply (erule natE, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   748
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   749
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   750
lemma nth_append [rule_format]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   751
  "xs:list(A) ==> 
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   752
   \<forall>n \<in> nat. nth(n, xs @ ys) = (if n < length(xs) then nth(n,xs)
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   753
                                else nth(n #- length(xs), ys))"
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   754
apply (induct_tac "xs", simp, clarify) 
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   755
apply (erule natE, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   756
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   757
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   758
lemma set_of_list_conv_nth:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   759
    "xs:list(A)
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   760
     ==> set_of_list(xs) = {x:A. EX i:nat. i<length(xs) & x = nth(i,xs)}"
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   761
apply (induct_tac "xs", simp_all)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   762
apply (rule equalityI, auto)
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   763
apply (rule_tac x = 0 in bexI, auto)
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   764
apply (erule natE, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   765
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   766
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   767
(* Other theorems about lists *)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   768
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   769
lemma nth_take_lemma [rule_format]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   770
 "k:nat ==>
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   771
  \<forall>xs \<in> list(A). (\<forall>ys \<in> list(A). k le length(xs) --> k le length(ys) -->
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   772
      (\<forall>i \<in> nat. i<k --> nth(i,xs) = nth(i,ys))--> take(k,xs) = take(k,ys))"
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   773
apply (induct_tac "k")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   774
apply (simp_all (no_asm_simp) add: lt_succ_eq_0_disj all_conj_distrib)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   775
apply clarify
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   776
(*Both lists are non-empty*)
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   777
apply (erule_tac a=xs in list.cases, simp) 
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   778
apply (erule_tac a=ys in list.cases, clarify) 
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   779
apply (simp (no_asm_use) )
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   780
apply clarify
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   781
apply (simp (no_asm_simp))
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   782
apply (rule conjI, force)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   783
apply (rename_tac y ys z zs) 
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   784
apply (drule_tac x = zs and x1 = ys in bspec [THEN bspec], auto)   
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   785
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   786
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   787
lemma nth_equalityI [rule_format]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   788
     "[| xs:list(A); ys:list(A); length(xs) = length(ys);
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   789
         \<forall>i \<in> nat. i < length(xs) --> nth(i,xs) = nth(i,ys) |]
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   790
      ==> xs = ys"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   791
apply (subgoal_tac "length (xs) le length (ys) ")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   792
apply (cut_tac k="length(xs)" and xs=xs and ys=ys in nth_take_lemma) 
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   793
apply (simp_all add: take_all)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   794
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   795
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   796
(*The famous take-lemma*)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   797
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   798
lemma take_equalityI [rule_format]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   799
    "[| xs:list(A); ys:list(A); (\<forall>i \<in> nat. take(i, xs) = take(i,ys)) |] 
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   800
     ==> xs = ys"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   801
apply (case_tac "length (xs) le length (ys) ")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   802
apply (drule_tac x = "length (ys) " in bspec)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   803
apply (drule_tac [3] not_lt_imp_le)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   804
apply (subgoal_tac [5] "length (ys) le length (xs) ")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   805
apply (rule_tac [6] j = "succ (length (ys))" in le_trans)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   806
apply (rule_tac [6] leI)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   807
apply (drule_tac [5] x = "length (xs) " in bspec)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   808
apply (simp_all add: take_all)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   809
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   810
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   811
lemma nth_drop [rule_format]:
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   812
  "n:nat ==> \<forall>i \<in> nat. \<forall>xs \<in> list(A). nth(i, drop(n, xs)) = nth(n #+ i, xs)"
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   813
apply (induct_tac "n", simp_all, clarify)
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   814
apply (erule list.cases, auto)  
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   815
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   816
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   817
subsection{*The function zip*}
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   818
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   819
text{*Crafty definition to eliminate a type argument*}
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   820
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   821
consts
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   822
  zip_aux        :: "[i,i]=>i"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   823
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   824
primrec (*explicit lambda is required because both arguments of "un" vary*)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   825
  "zip_aux(B,[]) =
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   826
     (\<lambda>ys \<in> list(B). list_case([], %y l. [], ys))"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   827
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   828
  "zip_aux(B,Cons(x,l)) =
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   829
     (\<lambda>ys \<in> list(B).
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   830
        list_case(Nil, %y zs. Cons(<x,y>, zip_aux(B,l)`zs), ys))"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   831
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   832
constdefs
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   833
  zip :: "[i, i]=>i"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   834
   "zip(xs, ys) == zip_aux(set_of_list(ys),xs)`ys"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   835
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   836
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   837
(* zip equations *)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   838
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   839
lemma list_on_set_of_list: "xs \<in> list(A) ==> xs \<in> list(set_of_list(xs))"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   840
apply (induct_tac xs, simp_all) 
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   841
apply (blast intro: list_mono [THEN subsetD]) 
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   842
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   843
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   844
lemma zip_Nil [simp]: "ys:list(A) ==> zip(Nil, ys)=Nil"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   845
apply (simp add: zip_def list_on_set_of_list [of _ A])
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   846
apply (erule list.cases, simp_all) 
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   847
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   848
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   849
lemma zip_Nil2 [simp]: "xs:list(A) ==> zip(xs, Nil)=Nil"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   850
apply (simp add: zip_def list_on_set_of_list [of _ A])
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   851
apply (erule list.cases, simp_all) 
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   852
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   853
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   854
lemma zip_aux_unique [rule_format]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   855
     "[|B<=C;  xs \<in> list(A)|] 
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   856
      ==> \<forall>ys \<in> list(B). zip_aux(C,xs) ` ys = zip_aux(B,xs) ` ys"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   857
apply (induct_tac xs) 
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   858
 apply simp_all 
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   859
 apply (blast intro: list_mono [THEN subsetD], clarify) 
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   860
apply (erule_tac a=ys in list.cases, auto) 
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   861
apply (blast intro: list_mono [THEN subsetD]) 
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   862
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   863
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   864
lemma zip_Cons_Cons [simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   865
     "[| xs:list(A); ys:list(B); x:A; y:B |] ==>
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   866
      zip(Cons(x,xs), Cons(y, ys)) = Cons(<x,y>, zip(xs, ys))"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   867
apply (simp add: zip_def, auto) 
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   868
apply (rule zip_aux_unique, auto) 
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   869
apply (simp add: list_on_set_of_list [of _ B])
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   870
apply (blast intro: list_on_set_of_list list_mono [THEN subsetD]) 
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   871
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   872
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   873
lemma zip_type [rule_format,simp,TC]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   874
     "xs:list(A) ==> \<forall>ys \<in> list(B). zip(xs, ys):list(A*B)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   875
apply (induct_tac "xs")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   876
apply (simp (no_asm))
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   877
apply clarify
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   878
apply (erule_tac a = ys in list.cases, auto)
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   879
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   880
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   881
(* zip length *)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   882
lemma length_zip [rule_format,simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   883
     "xs:list(A) ==> \<forall>ys \<in> list(B). length(zip(xs,ys)) =
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   884
                                     min(length(xs), length(ys))"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   885
apply (unfold min_def)
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   886
apply (induct_tac "xs", simp_all, clarify) 
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13327
diff changeset
   887
apply (erule_tac a = ys in list.cases, auto)
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   888
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   889
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   890
lemma zip_append1 [rule_format]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   891
 "[| ys:list(A); zs:list(B) |] ==>
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   892
  \<forall>xs \<in> list(A). zip(xs @ ys, zs) = 
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   893
                 zip(xs, take(length(xs), zs)) @ zip(ys, drop(length(xs),zs))"
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   894
apply (induct_tac "zs", force, clarify) 
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   895
apply (erule_tac a = xs in list.cases, simp_all) 
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   896
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   897
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   898
lemma zip_append2 [rule_format]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   899
 "[| xs:list(A); zs:list(B) |] ==> \<forall>ys \<in> list(B). zip(xs, ys@zs) =
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   900
       zip(take(length(ys), xs), ys) @ zip(drop(length(ys), xs), zs)"
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   901
apply (induct_tac "xs", force, clarify) 
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   902
apply (erule_tac a = ys in list.cases, auto)
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   903
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   904
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   905
lemma zip_append [simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   906
 "[| length(xs) = length(us); length(ys) = length(vs);
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   907
     xs:list(A); us:list(B); ys:list(A); vs:list(B) |] 
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   908
  ==> zip(xs@ys,us@vs) = zip(xs, us) @ zip(ys, vs)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   909
by (simp (no_asm_simp) add: zip_append1 drop_append diff_self_eq_0)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   910
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   911
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   912
lemma zip_rev [rule_format,simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   913
 "ys:list(B) ==> \<forall>xs \<in> list(A).
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   914
    length(xs) = length(ys) --> zip(rev(xs), rev(ys)) = rev(zip(xs, ys))"
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   915
apply (induct_tac "ys", force, clarify) 
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   916
apply (erule_tac a = xs in list.cases)
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   917
apply (auto simp add: length_rev)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   918
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   919
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   920
lemma nth_zip [rule_format,simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   921
   "ys:list(B) ==> \<forall>i \<in> nat. \<forall>xs \<in> list(A).
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   922
                    i < length(xs) --> i < length(ys) -->
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   923
                    nth(i,zip(xs, ys)) = <nth(i,xs),nth(i, ys)>"
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   924
apply (induct_tac "ys", force, clarify) 
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   925
apply (erule_tac a = xs in list.cases, simp) 
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   926
apply (auto elim: natE)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   927
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   928
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   929
lemma set_of_list_zip [rule_format]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   930
     "[| xs:list(A); ys:list(B); i:nat |]
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   931
      ==> set_of_list(zip(xs, ys)) =
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   932
          {<x, y>:A*B. EX i:nat. i < min(length(xs), length(ys))
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   933
          & x = nth(i, xs) & y = nth(i, ys)}"
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   934
by (force intro!: Collect_cong simp add: lt_min_iff set_of_list_conv_nth)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   935
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   936
(** list_update **)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   937
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   938
lemma list_update_Nil [simp]: "i:nat ==>list_update(Nil, i, v) = Nil"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   939
by (unfold list_update_def, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   940
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   941
lemma list_update_Cons_0 [simp]: "list_update(Cons(x, xs), 0, v)= Cons(v, xs)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   942
by (unfold list_update_def, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   943
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   944
lemma list_update_Cons_succ [simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   945
  "n:nat ==>
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   946
    list_update(Cons(x, xs), succ(n), v)= Cons(x, list_update(xs, n, v))"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   947
apply (unfold list_update_def, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   948
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   949
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   950
lemma list_update_type [rule_format,simp,TC]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   951
     "[| xs:list(A); v:A |] ==> \<forall>n \<in> nat. list_update(xs, n, v):list(A)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   952
apply (induct_tac "xs")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   953
apply (simp (no_asm))
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   954
apply clarify
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   955
apply (erule natE, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   956
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   957
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   958
lemma length_list_update [rule_format,simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   959
     "xs:list(A) ==> \<forall>i \<in> nat. length(list_update(xs, i, v))=length(xs)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   960
apply (induct_tac "xs")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   961
apply (simp (no_asm))
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   962
apply clarify
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   963
apply (erule natE, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   964
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   965
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   966
lemma nth_list_update [rule_format]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   967
     "[| xs:list(A) |] ==> \<forall>i \<in> nat. \<forall>j \<in> nat. i < length(xs)  -->
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   968
         nth(j, list_update(xs, i, x)) = (if i=j then x else nth(j, xs))"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   969
apply (induct_tac "xs")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   970
 apply simp_all
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   971
apply clarify
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   972
apply (rename_tac i j) 
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   973
apply (erule_tac n=i in natE) 
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   974
apply (erule_tac [2] n=j in natE)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   975
apply (erule_tac n=j in natE, simp_all, force) 
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   976
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   977
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   978
lemma nth_list_update_eq [simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   979
     "[| i < length(xs); xs:list(A) |] ==> nth(i, list_update(xs, i,x)) = x"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   980
by (simp (no_asm_simp) add: lt_length_in_nat nth_list_update)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   981
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   982
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   983
lemma nth_list_update_neq [rule_format,simp]:
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   984
  "xs:list(A) ==> 
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
   985
     \<forall>i \<in> nat. \<forall>j \<in> nat. i ~= j --> nth(j, list_update(xs,i,x)) = nth(j,xs)"
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   986
apply (induct_tac "xs")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   987
 apply (simp (no_asm))
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   988
apply clarify
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   989
apply (erule natE)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   990
apply (erule_tac [2] natE, simp_all) 
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   991
apply (erule natE, simp_all)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   992
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   993
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   994
lemma list_update_overwrite [rule_format,simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   995
     "xs:list(A) ==> \<forall>i \<in> nat. i < length(xs)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   996
   --> list_update(list_update(xs, i, x), i, y) = list_update(xs, i,y)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   997
apply (induct_tac "xs")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   998
 apply (simp (no_asm))
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
   999
apply clarify
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1000
apply (erule natE, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1001
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1002
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1003
lemma list_update_same_conv [rule_format]:
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1004
     "xs:list(A) ==> 
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1005
      \<forall>i \<in> nat. i < length(xs) --> 
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1006
                 (list_update(xs, i, x) = xs) <-> (nth(i, xs) = x)"
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1007
apply (induct_tac "xs")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1008
 apply (simp (no_asm))
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1009
apply clarify
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1010
apply (erule natE, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1011
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1012
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1013
lemma update_zip [rule_format]:
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1014
     "ys:list(B) ==> 
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1015
      \<forall>i \<in> nat. \<forall>xy \<in> A*B. \<forall>xs \<in> list(A).
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1016
        length(xs) = length(ys) -->
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1017
        list_update(zip(xs, ys), i, xy) = zip(list_update(xs, i, fst(xy)),
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1018
                                              list_update(ys, i, snd(xy)))"
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1019
apply (induct_tac "ys")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1020
 apply auto
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1021
apply (erule_tac a = xs in list.cases)
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1022
apply (auto elim: natE)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1023
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1024
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1025
lemma set_update_subset_cons [rule_format]:
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13327
diff changeset
  1026
  "xs:list(A) ==> 
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13327
diff changeset
  1027
   \<forall>i \<in> nat. set_of_list(list_update(xs, i, x)) <= cons(x, set_of_list(xs))"
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1028
apply (induct_tac "xs")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1029
 apply simp
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1030
apply (rule ballI)
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1031
apply (erule natE, simp_all, auto)
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1032
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1033
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1034
lemma set_of_list_update_subsetI:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1035
     "[| set_of_list(xs) <= A; xs:list(A); x:A; i:nat|]
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1036
   ==> set_of_list(list_update(xs, i,x)) <= A"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1037
apply (rule subset_trans)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1038
apply (rule set_update_subset_cons, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1039
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1040
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1041
(** upt **)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1042
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1043
lemma upt_rec:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1044
     "j:nat ==> upt(i,j) = (if i<j then Cons(i, upt(succ(i), j)) else Nil)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1045
apply (induct_tac "j", auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1046
apply (drule not_lt_imp_le)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1047
apply (auto simp: lt_Ord intro: le_anti_sym)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1048
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1049
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1050
lemma upt_conv_Nil [simp]: "[| j le i; j:nat |] ==> upt(i,j) = Nil"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1051
apply (subst upt_rec, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1052
apply (auto simp add: le_iff)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1053
apply (drule lt_asym [THEN notE], auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1054
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1055
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1056
(*Only needed if upt_Suc is deleted from the simpset*)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1057
lemma upt_succ_append:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1058
     "[| i le j; j:nat |] ==> upt(i,succ(j)) = upt(i, j)@[j]"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1059
by simp
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1060
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1061
lemma upt_conv_Cons:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1062
     "[| i<j; j:nat |]  ==> upt(i,j) = Cons(i,upt(succ(i),j))"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1063
apply (rule trans)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1064
apply (rule upt_rec, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1065
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1066
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1067
lemma upt_type [simp,TC]: "j:nat ==> upt(i,j):list(nat)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1068
by (induct_tac "j", auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1069
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1070
(*LOOPS as a simprule, since j<=j*)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1071
lemma upt_add_eq_append:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1072
     "[| i le j; j:nat; k:nat |] ==> upt(i, j #+k) = upt(i,j)@upt(j,j#+k)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1073
apply (induct_tac "k")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1074
apply (auto simp add: app_assoc app_type)
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1075
apply (rule_tac j = j in le_trans, auto)
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1076
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1077
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1078
lemma length_upt [simp]: "[| i:nat; j:nat |] ==>length(upt(i,j)) = j #- i"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1079
apply (induct_tac "j")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1080
apply (rule_tac [2] sym)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1081
apply (auto dest!: not_lt_imp_le simp add: length_app diff_succ diff_is_0_iff)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1082
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1083
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1084
lemma nth_upt [rule_format,simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1085
     "[| i:nat; j:nat; k:nat |] ==> i #+ k < j --> nth(k, upt(i,j)) = i #+ k"
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1086
apply (induct_tac "j", simp)
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1087
apply (simp add: nth_append le_iff split add: nat_diff_split)
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1088
apply (auto dest!: not_lt_imp_le 
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1089
            simp add: nth_append diff_self_eq_0 less_diff_conv add_commute)
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1090
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1091
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1092
lemma take_upt [rule_format,simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1093
     "[| m:nat; n:nat |] ==>
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1094
         \<forall>i \<in> nat. i #+ m le n --> take(m, upt(i,n)) = upt(i,i#+m)"
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1095
apply (induct_tac "m")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1096
apply (simp (no_asm_simp) add: take_0)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1097
apply clarify
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1098
apply (subst upt_rec, simp) 
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1099
apply (rule sym)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1100
apply (subst upt_rec, simp) 
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1101
apply (simp_all del: upt.simps)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1102
apply (rule_tac j = "succ (i #+ x) " in lt_trans2)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1103
apply auto
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1104
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1105
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1106
lemma map_succ_upt:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1107
     "[| m:nat; n:nat |] ==> map(succ, upt(m,n))= upt(succ(m), succ(n))"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1108
apply (induct_tac "n")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1109
apply (auto simp add: map_app_distrib)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1110
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1111
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1112
lemma nth_map [rule_format,simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1113
     "xs:list(A) ==>
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1114
      \<forall>n \<in> nat. n < length(xs) --> nth(n, map(f, xs)) = f(nth(n, xs))"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1115
apply (induct_tac "xs", simp)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1116
apply (rule ballI)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1117
apply (induct_tac "n", auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1118
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1119
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1120
lemma nth_map_upt [rule_format]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1121
     "[| m:nat; n:nat |] ==>
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1122
      \<forall>i \<in> nat. i < n #- m --> nth(i, map(f, upt(m,n))) = f(m #+ i)"
13784
b9f6154427a4 tidying (by script)
paulson
parents: 13615
diff changeset
  1123
apply (rule_tac n = m and m = n in diff_induct, typecheck, simp, simp) 
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1124
apply (subst map_succ_upt [symmetric], simp_all, clarify) 
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13327
diff changeset
  1125
apply (subgoal_tac "i < length (upt (0, x))")
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1126
 prefer 2 
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1127
 apply (simp add: less_diff_conv) 
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13327
diff changeset
  1128
 apply (rule_tac j = "succ (i #+ y) " in lt_trans2)
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1129
  apply simp 
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1130
 apply simp 
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13327
diff changeset
  1131
apply (subgoal_tac "i < length (upt (y, x))")
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1132
 apply (simp_all add: add_commute less_diff_conv)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1133
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1134
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1135
(** sublist (a generalization of nth to sets) **)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1136
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1137
constdefs
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1138
  sublist :: "[i, i] => i"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1139
    "sublist(xs, A)==
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1140
     map(fst, (filter(%p. snd(p): A, zip(xs, upt(0,length(xs))))))"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1141
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1142
lemma sublist_0 [simp]: "xs:list(A) ==>sublist(xs, 0) =Nil"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1143
by (unfold sublist_def, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1144
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1145
lemma sublist_Nil [simp]: "sublist(Nil, A) = Nil"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1146
by (unfold sublist_def, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1147
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1148
lemma sublist_shift_lemma:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1149
 "[| xs:list(B); i:nat |] ==>
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1150
  map(fst, filter(%p. snd(p):A, zip(xs, upt(i,i #+ length(xs))))) =
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1151
  map(fst, filter(%p. snd(p):nat & snd(p) #+ i:A, zip(xs,upt(0,length(xs)))))"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1152
apply (erule list_append_induct)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1153
apply (simp (no_asm_simp))
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1154
apply (auto simp add: add_commute length_app filter_append map_app_distrib)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1155
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1156
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1157
lemma sublist_type [simp,TC]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1158
     "xs:list(B) ==> sublist(xs, A):list(B)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1159
apply (unfold sublist_def)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1160
apply (induct_tac "xs")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1161
apply (auto simp add: filter_append map_app_distrib)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1162
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1163
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1164
lemma upt_add_eq_append2:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1165
     "[| i:nat; j:nat |] ==> upt(0, i #+ j) = upt(0, i) @ upt(i, i #+ j)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1166
by (simp add: upt_add_eq_append [of 0] nat_0_le)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1167
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1168
lemma sublist_append:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1169
 "[| xs:list(B); ys:list(B)  |] ==>
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1170
  sublist(xs@ys, A) = sublist(xs, A) @ sublist(ys, {j:nat. j #+ length(xs): A})"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1171
apply (unfold sublist_def)
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1172
apply (erule_tac l = ys in list_append_induct, simp)
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1173
apply (simp (no_asm_simp) add: upt_add_eq_append2 app_assoc [symmetric])
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1174
apply (auto simp add: sublist_shift_lemma length_type map_app_distrib app_assoc)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1175
apply (simp_all add: add_commute)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1176
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1177
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1178
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1179
lemma sublist_Cons:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1180
     "[| xs:list(B); x:B |] ==>
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1181
      sublist(Cons(x, xs), A) = 
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1182
      (if 0:A then [x] else []) @ sublist(xs, {j:nat. succ(j) : A})"
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1183
apply (erule_tac l = xs in list_append_induct)
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1184
apply (simp (no_asm_simp) add: sublist_def)  
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1185
apply (simp del: app_Cons add: app_Cons [symmetric] sublist_append, simp) 
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1186
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1187
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1188
lemma sublist_singleton [simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1189
     "sublist([x], A) = (if 0 : A then [x] else [])"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1190
by (simp (no_asm) add: sublist_Cons)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1191
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1192
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1193
lemma sublist_upt_eq_take [simp]:
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1194
    "[| xs:list(A); n:nat |] ==> sublist(xs, less_than(n)) = take(n,xs)"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1195
apply (unfold less_than_def)
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1196
apply (erule_tac l = xs in list_append_induct, simp) 
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1197
apply (simp split add: nat_diff_split add: sublist_append, auto)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1198
apply (subgoal_tac "n #- length (y) = 0")
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1199
apply (simp (no_asm_simp))
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1200
apply (auto dest!: not_lt_imp_le simp add: diff_is_0_iff)
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1201
done
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1202
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1203
text{*Repetition of a List Element*}
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1204
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1205
consts   repeat :: "[i,i]=>i"
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1206
primrec
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1207
  "repeat(a,0) = []"
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1208
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1209
  "repeat(a,succ(n)) = Cons(a,repeat(a,n))"
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1210
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1211
lemma length_repeat: "n \<in> nat ==> length(repeat(a,n)) = n"
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1212
by (induct_tac n, auto)
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1213
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1214
lemma repeat_succ_app: "n \<in> nat ==> repeat(a,succ(n)) = repeat(a,n) @ [a]"
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1215
apply (induct_tac n)
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1216
apply (simp_all del: app_Cons add: app_Cons [symmetric])
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1217
done
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1218
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1219
lemma repeat_type [TC]: "[|a \<in> A; n \<in> nat|] ==> repeat(a,n) \<in> list(A)"
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1220
by (induct_tac n, auto)
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1221
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13356
diff changeset
  1222
13327
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1223
ML
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1224
{*
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1225
val ConsE = thm "ConsE";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1226
val Cons_iff = thm "Cons_iff";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1227
val Nil_Cons_iff = thm "Nil_Cons_iff";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1228
val list_unfold = thm "list_unfold";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1229
val list_mono = thm "list_mono";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1230
val list_univ = thm "list_univ";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1231
val list_subset_univ = thm "list_subset_univ";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1232
val list_into_univ = thm "list_into_univ";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1233
val list_case_type = thm "list_case_type";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1234
val tl_type = thm "tl_type";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1235
val drop_Nil = thm "drop_Nil";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1236
val drop_succ_Cons = thm "drop_succ_Cons";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1237
val drop_type = thm "drop_type";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1238
val list_rec_type = thm "list_rec_type";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1239
val map_type = thm "map_type";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1240
val map_type2 = thm "map_type2";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1241
val length_type = thm "length_type";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1242
val lt_length_in_nat = thm "lt_length_in_nat";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1243
val app_type = thm "app_type";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1244
val rev_type = thm "rev_type";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1245
val flat_type = thm "flat_type";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1246
val set_of_list_type = thm "set_of_list_type";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1247
val set_of_list_append = thm "set_of_list_append";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1248
val list_add_type = thm "list_add_type";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1249
val map_ident = thm "map_ident";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1250
val map_compose = thm "map_compose";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1251
val map_app_distrib = thm "map_app_distrib";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1252
val map_flat = thm "map_flat";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1253
val list_rec_map = thm "list_rec_map";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1254
val list_CollectD = thm "list_CollectD";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1255
val map_list_Collect = thm "map_list_Collect";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1256
val length_map = thm "length_map";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1257
val length_app = thm "length_app";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1258
val length_rev = thm "length_rev";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1259
val length_flat = thm "length_flat";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1260
val drop_length_Cons = thm "drop_length_Cons";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1261
val drop_length = thm "drop_length";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1262
val app_right_Nil = thm "app_right_Nil";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1263
val app_assoc = thm "app_assoc";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1264
val flat_app_distrib = thm "flat_app_distrib";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1265
val rev_map_distrib = thm "rev_map_distrib";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1266
val rev_app_distrib = thm "rev_app_distrib";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1267
val rev_rev_ident = thm "rev_rev_ident";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1268
val rev_flat = thm "rev_flat";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1269
val list_add_app = thm "list_add_app";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1270
val list_add_rev = thm "list_add_rev";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1271
val list_add_flat = thm "list_add_flat";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1272
val list_append_induct = thm "list_append_induct";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1273
val min_sym = thm "min_sym";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1274
val min_type = thm "min_type";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1275
val min_0 = thm "min_0";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1276
val min_02 = thm "min_02";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1277
val lt_min_iff = thm "lt_min_iff";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1278
val min_succ_succ = thm "min_succ_succ";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1279
val filter_append = thm "filter_append";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1280
val filter_type = thm "filter_type";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1281
val length_filter = thm "length_filter";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1282
val filter_is_subset = thm "filter_is_subset";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1283
val filter_False = thm "filter_False";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1284
val filter_True = thm "filter_True";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1285
val length_is_0_iff = thm "length_is_0_iff";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1286
val length_is_0_iff2 = thm "length_is_0_iff2";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1287
val length_tl = thm "length_tl";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1288
val length_greater_0_iff = thm "length_greater_0_iff";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1289
val length_succ_iff = thm "length_succ_iff";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1290
val append_is_Nil_iff = thm "append_is_Nil_iff";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1291
val append_is_Nil_iff2 = thm "append_is_Nil_iff2";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1292
val append_left_is_self_iff = thm "append_left_is_self_iff";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1293
val append_left_is_self_iff2 = thm "append_left_is_self_iff2";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1294
val append_left_is_Nil_iff = thm "append_left_is_Nil_iff";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1295
val append_left_is_Nil_iff2 = thm "append_left_is_Nil_iff2";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1296
val append_eq_append_iff = thm "append_eq_append_iff";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1297
val append_eq_append = thm "append_eq_append";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1298
val append_eq_append_iff2 = thm "append_eq_append_iff2";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1299
val append_self_iff = thm "append_self_iff";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1300
val append_self_iff2 = thm "append_self_iff2";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1301
val append1_eq_iff = thm "append1_eq_iff";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1302
val append_right_is_self_iff = thm "append_right_is_self_iff";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1303
val append_right_is_self_iff2 = thm "append_right_is_self_iff2";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1304
val hd_append = thm "hd_append";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1305
val tl_append = thm "tl_append";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1306
val rev_is_Nil_iff = thm "rev_is_Nil_iff";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1307
val Nil_is_rev_iff = thm "Nil_is_rev_iff";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1308
val rev_is_rev_iff = thm "rev_is_rev_iff";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1309
val rev_list_elim = thm "rev_list_elim";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1310
val length_drop = thm "length_drop";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1311
val drop_all = thm "drop_all";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1312
val drop_append = thm "drop_append";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1313
val drop_drop = thm "drop_drop";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1314
val take_0 = thm "take_0";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1315
val take_succ_Cons = thm "take_succ_Cons";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1316
val take_Nil = thm "take_Nil";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1317
val take_all = thm "take_all";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1318
val take_type = thm "take_type";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1319
val take_append = thm "take_append";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1320
val take_take = thm "take_take";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1321
val nth_0 = thm "nth_0";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1322
val nth_Cons = thm "nth_Cons";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1323
val nth_type = thm "nth_type";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1324
val nth_append = thm "nth_append";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1325
val set_of_list_conv_nth = thm "set_of_list_conv_nth";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1326
val nth_take_lemma = thm "nth_take_lemma";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1327
val nth_equalityI = thm "nth_equalityI";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1328
val take_equalityI = thm "take_equalityI";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1329
val nth_drop = thm "nth_drop";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1330
val list_on_set_of_list = thm "list_on_set_of_list";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1331
val zip_Nil = thm "zip_Nil";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1332
val zip_Nil2 = thm "zip_Nil2";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1333
val zip_Cons_Cons = thm "zip_Cons_Cons";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1334
val zip_type = thm "zip_type";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1335
val length_zip = thm "length_zip";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1336
val zip_append1 = thm "zip_append1";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1337
val zip_append2 = thm "zip_append2";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1338
val zip_append = thm "zip_append";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1339
val zip_rev = thm "zip_rev";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1340
val nth_zip = thm "nth_zip";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1341
val set_of_list_zip = thm "set_of_list_zip";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1342
val list_update_Nil = thm "list_update_Nil";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1343
val list_update_Cons_0 = thm "list_update_Cons_0";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1344
val list_update_Cons_succ = thm "list_update_Cons_succ";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1345
val list_update_type = thm "list_update_type";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1346
val length_list_update = thm "length_list_update";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1347
val nth_list_update = thm "nth_list_update";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1348
val nth_list_update_eq = thm "nth_list_update_eq";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1349
val nth_list_update_neq = thm "nth_list_update_neq";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1350
val list_update_overwrite = thm "list_update_overwrite";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1351
val list_update_same_conv = thm "list_update_same_conv";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1352
val update_zip = thm "update_zip";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1353
val set_update_subset_cons = thm "set_update_subset_cons";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1354
val set_of_list_update_subsetI = thm "set_of_list_update_subsetI";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1355
val upt_rec = thm "upt_rec";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1356
val upt_conv_Nil = thm "upt_conv_Nil";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1357
val upt_succ_append = thm "upt_succ_append";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1358
val upt_conv_Cons = thm "upt_conv_Cons";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1359
val upt_type = thm "upt_type";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1360
val upt_add_eq_append = thm "upt_add_eq_append";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1361
val length_upt = thm "length_upt";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1362
val nth_upt = thm "nth_upt";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1363
val take_upt = thm "take_upt";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1364
val map_succ_upt = thm "map_succ_upt";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1365
val nth_map = thm "nth_map";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1366
val nth_map_upt = thm "nth_map_upt";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1367
val sublist_0 = thm "sublist_0";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1368
val sublist_Nil = thm "sublist_Nil";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1369
val sublist_shift_lemma = thm "sublist_shift_lemma";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1370
val sublist_type = thm "sublist_type";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1371
val upt_add_eq_append2 = thm "upt_add_eq_append2";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1372
val sublist_append = thm "sublist_append";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1373
val sublist_Cons = thm "sublist_Cons";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1374
val sublist_singleton = thm "sublist_singleton";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1375
val sublist_upt_eq_take = thm "sublist_upt_eq_take";
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1376
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1377
structure list =
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1378
struct
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1379
val induct = thm "list.induct"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1380
val elim   = thm "list.cases"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1381
val intrs  = thms "list.intros"
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1382
end;  
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1383
*}
be7105a066d3 converted List to new-style
paulson
parents: 12789
diff changeset
  1384
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
  1385
end