src/HOL/List.thy
author nipkow
Mon, 10 Jan 2000 16:06:43 +0100
changeset 8115 c802042066e8
parent 8000 acafa0f15131
child 8490 6e0f23304061
permissions -rw-r--r--
Forgot to "call" MicroJava in makefile. Added list_all2 to List.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     1
(*  Title:      HOL/List.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     3
    Author:     Tobias Nipkow
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     4
    Copyright   1994 TU Muenchen
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     5
2512
0231e4f467f2 Got rid of Alls in List.
nipkow
parents: 2369
diff changeset
     6
The datatype of finite lists.
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     7
*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     8
7032
d6efb3b8e669 NatBin: binary arithmetic for the naturals
paulson
parents: 6408
diff changeset
     9
List = Datatype + WF_Rel + NatBin +
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    10
7224
e41e64476f9b 'a list: Nil, Cons;
wenzelm
parents: 7032
diff changeset
    11
datatype 'a list = Nil ("[]") | Cons 'a ('a list) (infixr "#" 65)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    12
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    13
consts
1908
55d8e38262a8 Renamed setOfList to set_of_list
paulson
parents: 1898
diff changeset
    14
  "@"         :: ['a list, 'a list] => 'a list            (infixr 65)
55d8e38262a8 Renamed setOfList to set_of_list
paulson
parents: 1898
diff changeset
    15
  filter      :: ['a => bool, 'a list] => 'a list
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents: 2512
diff changeset
    16
  concat      :: 'a list list => 'a list
1908
55d8e38262a8 Renamed setOfList to set_of_list
paulson
parents: 1898
diff changeset
    17
  foldl       :: [['b,'a] => 'b, 'b, 'a list] => 'b
8000
acafa0f15131 added foldr
paulson
parents: 7224
diff changeset
    18
  foldr       :: [['a,'b] => 'b, 'a list, 'b] => 'b
3896
ee8ebb74ec00 Various new lemmas. Improved conversion of equations to rewrite rules:
nipkow
parents: 3842
diff changeset
    19
  hd, last    :: 'a list => 'a
3465
e85c24717cad set_of_list -> set
nipkow
parents: 3437
diff changeset
    20
  set         :: 'a list => 'a set
5518
654ead0ba4f7 re-added mem and list_all
oheimb
parents: 5443
diff changeset
    21
  list_all    :: ('a => bool) => ('a list => bool)
8115
c802042066e8 Forgot to "call" MicroJava in makefile.
nipkow
parents: 8000
diff changeset
    22
  list_all2   :: ('a => 'b => bool) => 'a list => 'b list => bool
1908
55d8e38262a8 Renamed setOfList to set_of_list
paulson
parents: 1898
diff changeset
    23
  map         :: ('a=>'b) => ('a list => 'b list)
5518
654ead0ba4f7 re-added mem and list_all
oheimb
parents: 5443
diff changeset
    24
  mem         :: ['a, 'a list] => bool                    (infixl 55)
4502
337c073de95e nth -> !
nipkow
parents: 4151
diff changeset
    25
  nth         :: ['a list, nat] => 'a			  (infixl "!" 100)
5077
71043526295f * HOL/List: new function list_update written xs[i:=v] that updates the i-th
nipkow
parents: 4643
diff changeset
    26
  list_update :: 'a list => nat => 'a => 'a list
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents: 2512
diff changeset
    27
  take, drop  :: [nat, 'a list] => 'a list
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents: 2512
diff changeset
    28
  takeWhile,
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents: 2512
diff changeset
    29
  dropWhile   :: ('a => bool) => 'a list => 'a list
3896
ee8ebb74ec00 Various new lemmas. Improved conversion of equations to rewrite rules:
nipkow
parents: 3842
diff changeset
    30
  tl, butlast :: 'a list => 'a list
1908
55d8e38262a8 Renamed setOfList to set_of_list
paulson
parents: 1898
diff changeset
    31
  rev         :: 'a list => 'a list
4132
daff3c9987cc added zip and nodup
oheimb
parents: 3896
diff changeset
    32
  zip	      :: "'a list => 'b list => ('a * 'b) list"
5427
26c9a7c0b36b Arith: less_diff_conv
nipkow
parents: 5425
diff changeset
    33
  upt         :: nat => nat => nat list                   ("(1[_../_'(])")
4605
579e0ef2df6b Added `remdups'
nipkow
parents: 4502
diff changeset
    34
  remdups     :: 'a list => 'a list
579e0ef2df6b Added `remdups'
nipkow
parents: 4502
diff changeset
    35
  nodups      :: "'a list => bool"
3589
244daa75f890 Added function `replicate' and lemmas map_cong and set_replicate.
nipkow
parents: 3584
diff changeset
    36
  replicate   :: nat => 'a => 'a list
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    37
5077
71043526295f * HOL/List: new function list_update written xs[i:=v] that updates the i-th
nipkow
parents: 4643
diff changeset
    38
nonterminals
71043526295f * HOL/List: new function list_update written xs[i:=v] that updates the i-th
nipkow
parents: 4643
diff changeset
    39
  lupdbinds  lupdbind
71043526295f * HOL/List: new function list_update written xs[i:=v] that updates the i-th
nipkow
parents: 4643
diff changeset
    40
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    41
syntax
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    42
  (* list Enumeration *)
2262
c7ee913746fd added symbols syntax;
wenzelm
parents: 1908
diff changeset
    43
  "@list"     :: args => 'a list                          ("[(_)]")
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    44
2512
0231e4f467f2 Got rid of Alls in List.
nipkow
parents: 2369
diff changeset
    45
  (* Special syntax for filter *)
5295
25fb5156e0d9 replaced idt by pttrn in @filter
oheimb
parents: 5281
diff changeset
    46
  "@filter"   :: [pttrn, 'a list, bool] => 'a list        ("(1[_:_ ./ _])")
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    47
5077
71043526295f * HOL/List: new function list_update written xs[i:=v] that updates the i-th
nipkow
parents: 4643
diff changeset
    48
  (* list update *)
71043526295f * HOL/List: new function list_update written xs[i:=v] that updates the i-th
nipkow
parents: 4643
diff changeset
    49
  "_lupdbind"      :: ['a, 'a] => lupdbind            ("(2_ :=/ _)")
71043526295f * HOL/List: new function list_update written xs[i:=v] that updates the i-th
nipkow
parents: 4643
diff changeset
    50
  ""               :: lupdbind => lupdbinds           ("_")
71043526295f * HOL/List: new function list_update written xs[i:=v] that updates the i-th
nipkow
parents: 4643
diff changeset
    51
  "_lupdbinds"     :: [lupdbind, lupdbinds] => lupdbinds ("_,/ _")
71043526295f * HOL/List: new function list_update written xs[i:=v] that updates the i-th
nipkow
parents: 4643
diff changeset
    52
  "_LUpdate"       :: ['a, lupdbinds] => 'a           ("_/[(_)]" [900,0] 900)
71043526295f * HOL/List: new function list_update written xs[i:=v] that updates the i-th
nipkow
parents: 4643
diff changeset
    53
5427
26c9a7c0b36b Arith: less_diff_conv
nipkow
parents: 5425
diff changeset
    54
  upto        :: nat => nat => nat list                   ("(1[_../_])")
26c9a7c0b36b Arith: less_diff_conv
nipkow
parents: 5425
diff changeset
    55
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    56
translations
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    57
  "[x, xs]"     == "x#[xs]"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    58
  "[x]"         == "x#[]"
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3589
diff changeset
    59
  "[x:xs . P]"  == "filter (%x. P) xs"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    60
5077
71043526295f * HOL/List: new function list_update written xs[i:=v] that updates the i-th
nipkow
parents: 4643
diff changeset
    61
  "_LUpdate xs (_lupdbinds b bs)"  == "_LUpdate (_LUpdate xs b) bs"
71043526295f * HOL/List: new function list_update written xs[i:=v] that updates the i-th
nipkow
parents: 4643
diff changeset
    62
  "xs[i:=x]"                       == "list_update xs i x"
71043526295f * HOL/List: new function list_update written xs[i:=v] that updates the i-th
nipkow
parents: 4643
diff changeset
    63
5427
26c9a7c0b36b Arith: less_diff_conv
nipkow
parents: 5425
diff changeset
    64
  "[i..j]" == "[i..(Suc j)(]"
26c9a7c0b36b Arith: less_diff_conv
nipkow
parents: 5425
diff changeset
    65
26c9a7c0b36b Arith: less_diff_conv
nipkow
parents: 5425
diff changeset
    66
2262
c7ee913746fd added symbols syntax;
wenzelm
parents: 1908
diff changeset
    67
syntax (symbols)
5295
25fb5156e0d9 replaced idt by pttrn in @filter
oheimb
parents: 5281
diff changeset
    68
  "@filter"   :: [pttrn, 'a list, bool] => 'a list        ("(1[_\\<in>_ ./ _])")
2262
c7ee913746fd added symbols syntax;
wenzelm
parents: 1908
diff changeset
    69
c7ee913746fd added symbols syntax;
wenzelm
parents: 1908
diff changeset
    70
3342
ec3b55fcb165 New operator "lists" for formalizing sets of lists
paulson
parents: 3320
diff changeset
    71
consts
ec3b55fcb165 New operator "lists" for formalizing sets of lists
paulson
parents: 3320
diff changeset
    72
  lists        :: 'a set => 'a list set
ec3b55fcb165 New operator "lists" for formalizing sets of lists
paulson
parents: 3320
diff changeset
    73
ec3b55fcb165 New operator "lists" for formalizing sets of lists
paulson
parents: 3320
diff changeset
    74
  inductive "lists A"
ec3b55fcb165 New operator "lists" for formalizing sets of lists
paulson
parents: 3320
diff changeset
    75
  intrs
ec3b55fcb165 New operator "lists" for formalizing sets of lists
paulson
parents: 3320
diff changeset
    76
    Nil  "[]: lists A"
ec3b55fcb165 New operator "lists" for formalizing sets of lists
paulson
parents: 3320
diff changeset
    77
    Cons "[| a: A;  l: lists A |] ==> a#l : lists A"
ec3b55fcb165 New operator "lists" for formalizing sets of lists
paulson
parents: 3320
diff changeset
    78
ec3b55fcb165 New operator "lists" for formalizing sets of lists
paulson
parents: 3320
diff changeset
    79
3437
bea2faf1641d Replacing the primrec definition of "length" by a translation to the built-in
paulson
parents: 3401
diff changeset
    80
(*Function "size" is overloaded for all datatypes.  Users may refer to the
bea2faf1641d Replacing the primrec definition of "length" by a translation to the built-in
paulson
parents: 3401
diff changeset
    81
  list version as "length".*)
bea2faf1641d Replacing the primrec definition of "length" by a translation to the built-in
paulson
parents: 3401
diff changeset
    82
syntax   length :: 'a list => nat
3507
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
    83
translations  "length"  =>  "size:: _ list => nat"
3437
bea2faf1641d Replacing the primrec definition of "length" by a translation to the built-in
paulson
parents: 3401
diff changeset
    84
5183
89f162de39cf Adapted to new datatype package.
berghofe
parents: 5162
diff changeset
    85
primrec
3320
3a5e4930fb77 Added `arbitrary'
nipkow
parents: 3196
diff changeset
    86
  "hd([]) = arbitrary"
1898
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    87
  "hd(x#xs) = x"
5183
89f162de39cf Adapted to new datatype package.
berghofe
parents: 5162
diff changeset
    88
primrec
3896
ee8ebb74ec00 Various new lemmas. Improved conversion of equations to rewrite rules:
nipkow
parents: 3842
diff changeset
    89
  "tl([]) = []"
1898
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    90
  "tl(x#xs) = xs"
5183
89f162de39cf Adapted to new datatype package.
berghofe
parents: 5162
diff changeset
    91
primrec
3896
ee8ebb74ec00 Various new lemmas. Improved conversion of equations to rewrite rules:
nipkow
parents: 3842
diff changeset
    92
  "last [] = arbitrary"
ee8ebb74ec00 Various new lemmas. Improved conversion of equations to rewrite rules:
nipkow
parents: 3842
diff changeset
    93
  "last(x#xs) = (if xs=[] then x else last xs)"
5183
89f162de39cf Adapted to new datatype package.
berghofe
parents: 5162
diff changeset
    94
primrec
3896
ee8ebb74ec00 Various new lemmas. Improved conversion of equations to rewrite rules:
nipkow
parents: 3842
diff changeset
    95
  "butlast [] = []"
ee8ebb74ec00 Various new lemmas. Improved conversion of equations to rewrite rules:
nipkow
parents: 3842
diff changeset
    96
  "butlast(x#xs) = (if xs=[] then [] else x#butlast xs)"
5183
89f162de39cf Adapted to new datatype package.
berghofe
parents: 5162
diff changeset
    97
primrec
5518
654ead0ba4f7 re-added mem and list_all
oheimb
parents: 5443
diff changeset
    98
  "x mem [] = False"
654ead0ba4f7 re-added mem and list_all
oheimb
parents: 5443
diff changeset
    99
  "x mem (y#ys) = (if y=x then True else x mem ys)"
654ead0ba4f7 re-added mem and list_all
oheimb
parents: 5443
diff changeset
   100
primrec
3465
e85c24717cad set_of_list -> set
nipkow
parents: 3437
diff changeset
   101
  "set [] = {}"
e85c24717cad set_of_list -> set
nipkow
parents: 3437
diff changeset
   102
  "set (x#xs) = insert x (set xs)"
5183
89f162de39cf Adapted to new datatype package.
berghofe
parents: 5162
diff changeset
   103
primrec
5518
654ead0ba4f7 re-added mem and list_all
oheimb
parents: 5443
diff changeset
   104
  list_all_Nil  "list_all P [] = True"
654ead0ba4f7 re-added mem and list_all
oheimb
parents: 5443
diff changeset
   105
  list_all_Cons "list_all P (x#xs) = (P(x) & list_all P xs)"
654ead0ba4f7 re-added mem and list_all
oheimb
parents: 5443
diff changeset
   106
primrec
1898
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
   107
  "map f [] = []"
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
   108
  "map f (x#xs) = f(x)#map f xs"
5183
89f162de39cf Adapted to new datatype package.
berghofe
parents: 5162
diff changeset
   109
primrec
89f162de39cf Adapted to new datatype package.
berghofe
parents: 5162
diff changeset
   110
  append_Nil  "[]    @ys = ys"
89f162de39cf Adapted to new datatype package.
berghofe
parents: 5162
diff changeset
   111
  append_Cons "(x#xs)@ys = x#(xs@ys)"
89f162de39cf Adapted to new datatype package.
berghofe
parents: 5162
diff changeset
   112
primrec
1898
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
   113
  "rev([]) = []"
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
   114
  "rev(x#xs) = rev(xs) @ [x]"
5183
89f162de39cf Adapted to new datatype package.
berghofe
parents: 5162
diff changeset
   115
primrec
1898
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
   116
  "filter P [] = []"
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
   117
  "filter P (x#xs) = (if P x then x#filter P xs else filter P xs)"
5183
89f162de39cf Adapted to new datatype package.
berghofe
parents: 5162
diff changeset
   118
primrec
6141
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 5518
diff changeset
   119
  foldl_Nil  "foldl f a [] = a"
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 5518
diff changeset
   120
  foldl_Cons "foldl f a (x#xs) = foldl f (f a x) xs"
5183
89f162de39cf Adapted to new datatype package.
berghofe
parents: 5162
diff changeset
   121
primrec
8000
acafa0f15131 added foldr
paulson
parents: 7224
diff changeset
   122
  "foldr f [] a = a"
acafa0f15131 added foldr
paulson
parents: 7224
diff changeset
   123
  "foldr f (x#xs) a = f x (foldr f xs a)"
acafa0f15131 added foldr
paulson
parents: 7224
diff changeset
   124
primrec
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents: 2512
diff changeset
   125
  "concat([]) = []"
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents: 2512
diff changeset
   126
  "concat(x#xs) = x @ concat(xs)"
5183
89f162de39cf Adapted to new datatype package.
berghofe
parents: 5162
diff changeset
   127
primrec
1419
a6a034a47a71 defined take/drop by induction over list rather than nat.
nipkow
parents: 1370
diff changeset
   128
  drop_Nil  "drop n [] = []"
a6a034a47a71 defined take/drop by induction over list rather than nat.
nipkow
parents: 1370
diff changeset
   129
  drop_Cons "drop n (x#xs) = (case n of 0 => x#xs | Suc(m) => drop m xs)"
6408
5b443d6331ed new definition for nth.
pusch
parents: 6306
diff changeset
   130
  (* Warning: simpset does not contain this definition but separate theorems 
5b443d6331ed new definition for nth.
pusch
parents: 6306
diff changeset
   131
     for n=0 / n=Suc k*)
5183
89f162de39cf Adapted to new datatype package.
berghofe
parents: 5162
diff changeset
   132
primrec
1419
a6a034a47a71 defined take/drop by induction over list rather than nat.
nipkow
parents: 1370
diff changeset
   133
  take_Nil  "take n [] = []"
a6a034a47a71 defined take/drop by induction over list rather than nat.
nipkow
parents: 1370
diff changeset
   134
  take_Cons "take n (x#xs) = (case n of 0 => [] | Suc(m) => x # take m xs)"
6408
5b443d6331ed new definition for nth.
pusch
parents: 6306
diff changeset
   135
  (* Warning: simpset does not contain this definition but separate theorems 
5b443d6331ed new definition for nth.
pusch
parents: 6306
diff changeset
   136
     for n=0 / n=Suc k*)
5b443d6331ed new definition for nth.
pusch
parents: 6306
diff changeset
   137
primrec 
5b443d6331ed new definition for nth.
pusch
parents: 6306
diff changeset
   138
  nth_Cons  "(x#xs)!n = (case n of 0 => x | (Suc k) => xs!k)"
5b443d6331ed new definition for nth.
pusch
parents: 6306
diff changeset
   139
  (* Warning: simpset does not contain this definition but separate theorems 
5b443d6331ed new definition for nth.
pusch
parents: 6306
diff changeset
   140
     for n=0 / n=Suc k*)
5183
89f162de39cf Adapted to new datatype package.
berghofe
parents: 5162
diff changeset
   141
primrec
5077
71043526295f * HOL/List: new function list_update written xs[i:=v] that updates the i-th
nipkow
parents: 4643
diff changeset
   142
 "    [][i:=v] = []"
71043526295f * HOL/List: new function list_update written xs[i:=v] that updates the i-th
nipkow
parents: 4643
diff changeset
   143
 "(x#xs)[i:=v] = (case i of 0     => v # xs 
71043526295f * HOL/List: new function list_update written xs[i:=v] that updates the i-th
nipkow
parents: 4643
diff changeset
   144
			  | Suc j => x # xs[j:=v])"
5183
89f162de39cf Adapted to new datatype package.
berghofe
parents: 5162
diff changeset
   145
primrec
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents: 2512
diff changeset
   146
  "takeWhile P [] = []"
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents: 2512
diff changeset
   147
  "takeWhile P (x#xs) = (if P x then x#takeWhile P xs else [])"
5183
89f162de39cf Adapted to new datatype package.
berghofe
parents: 5162
diff changeset
   148
primrec
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents: 2512
diff changeset
   149
  "dropWhile P [] = []"
3584
8f9ee0f79d9a Corected bug in def of dropWhile (also present in Haskell lib!)
nipkow
parents: 3507
diff changeset
   150
  "dropWhile P (x#xs) = (if P x then dropWhile P xs else x#xs)"
5183
89f162de39cf Adapted to new datatype package.
berghofe
parents: 5162
diff changeset
   151
primrec
4132
daff3c9987cc added zip and nodup
oheimb
parents: 3896
diff changeset
   152
  "zip xs []     = []"
6306
81e7fbf61db2 modified zip
nipkow
parents: 6141
diff changeset
   153
  "zip xs (y#ys) = (case xs of [] => [] | z#zs => (z,y)#zip zs ys)"
6408
5b443d6331ed new definition for nth.
pusch
parents: 6306
diff changeset
   154
  (* Warning: simpset does not contain this definition but separate theorems 
5b443d6331ed new definition for nth.
pusch
parents: 6306
diff changeset
   155
     for xs=[] / xs=z#zs *)
5427
26c9a7c0b36b Arith: less_diff_conv
nipkow
parents: 5425
diff changeset
   156
primrec
26c9a7c0b36b Arith: less_diff_conv
nipkow
parents: 5425
diff changeset
   157
  "[i..0(] = []"
26c9a7c0b36b Arith: less_diff_conv
nipkow
parents: 5425
diff changeset
   158
  "[i..(Suc j)(] = (if i <= j then [i..j(] @ [j] else [])"
5183
89f162de39cf Adapted to new datatype package.
berghofe
parents: 5162
diff changeset
   159
primrec
4605
579e0ef2df6b Added `remdups'
nipkow
parents: 4502
diff changeset
   160
  "nodups []     = True"
579e0ef2df6b Added `remdups'
nipkow
parents: 4502
diff changeset
   161
  "nodups (x#xs) = (x ~: set xs & nodups xs)"
5183
89f162de39cf Adapted to new datatype package.
berghofe
parents: 5162
diff changeset
   162
primrec
4605
579e0ef2df6b Added `remdups'
nipkow
parents: 4502
diff changeset
   163
  "remdups [] = []"
579e0ef2df6b Added `remdups'
nipkow
parents: 4502
diff changeset
   164
  "remdups (x#xs) = (if x : set xs then remdups xs else x # remdups xs)"
5183
89f162de39cf Adapted to new datatype package.
berghofe
parents: 5162
diff changeset
   165
primrec
5443
e2459d18ff47 changed constants mem and list_all to mere translations
oheimb
parents: 5427
diff changeset
   166
  replicate_0   "replicate  0      x = []"
5183
89f162de39cf Adapted to new datatype package.
berghofe
parents: 5162
diff changeset
   167
  replicate_Suc "replicate (Suc n) x = x # replicate n x"
8115
c802042066e8 Forgot to "call" MicroJava in makefile.
nipkow
parents: 8000
diff changeset
   168
defs
c802042066e8 Forgot to "call" MicroJava in makefile.
nipkow
parents: 8000
diff changeset
   169
 list_all2_def
c802042066e8 Forgot to "call" MicroJava in makefile.
nipkow
parents: 8000
diff changeset
   170
 "list_all2 P xs ys == length xs = length ys & (!(x,y):set(zip xs ys). P x y)"
c802042066e8 Forgot to "call" MicroJava in makefile.
nipkow
parents: 8000
diff changeset
   171
3196
c522bc46aea7 Added pred_list for TFL
paulson
parents: 2738
diff changeset
   172
6408
5b443d6331ed new definition for nth.
pusch
parents: 6306
diff changeset
   173
(** Lexicographic orderings on lists **)
5281
f4d16517b360 List now contains some lexicographic orderings.
nipkow
parents: 5183
diff changeset
   174
f4d16517b360 List now contains some lexicographic orderings.
nipkow
parents: 5183
diff changeset
   175
consts
f4d16517b360 List now contains some lexicographic orderings.
nipkow
parents: 5183
diff changeset
   176
 lexn :: "('a * 'a)set => nat => ('a list * 'a list)set"
f4d16517b360 List now contains some lexicographic orderings.
nipkow
parents: 5183
diff changeset
   177
primrec
f4d16517b360 List now contains some lexicographic orderings.
nipkow
parents: 5183
diff changeset
   178
"lexn r 0       = {}"
f4d16517b360 List now contains some lexicographic orderings.
nipkow
parents: 5183
diff changeset
   179
"lexn r (Suc n) = (prod_fun (split op#) (split op#) `` (r ** lexn r n)) Int
f4d16517b360 List now contains some lexicographic orderings.
nipkow
parents: 5183
diff changeset
   180
                  {(xs,ys). length xs = Suc n & length ys = Suc n}"
f4d16517b360 List now contains some lexicographic orderings.
nipkow
parents: 5183
diff changeset
   181
f4d16517b360 List now contains some lexicographic orderings.
nipkow
parents: 5183
diff changeset
   182
constdefs
f4d16517b360 List now contains some lexicographic orderings.
nipkow
parents: 5183
diff changeset
   183
 lex :: "('a * 'a)set => ('a list * 'a list)set"
f4d16517b360 List now contains some lexicographic orderings.
nipkow
parents: 5183
diff changeset
   184
"lex r == UN n. lexn r n"
f4d16517b360 List now contains some lexicographic orderings.
nipkow
parents: 5183
diff changeset
   185
f4d16517b360 List now contains some lexicographic orderings.
nipkow
parents: 5183
diff changeset
   186
 lexico :: "('a * 'a)set => ('a list * 'a list)set"
f4d16517b360 List now contains some lexicographic orderings.
nipkow
parents: 5183
diff changeset
   187
"lexico r == inv_image (less_than ** lex r) (%xs. (length xs, xs))"
f4d16517b360 List now contains some lexicographic orderings.
nipkow
parents: 5183
diff changeset
   188
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   189
end
3507
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
   190
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
   191
ML
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
   192
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
   193
local
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
   194
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
   195
(* translating size::list -> length *)
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
   196
4151
5c19cd418c33 adapted typed_print_translation;
wenzelm
parents: 4132
diff changeset
   197
fun size_tr' _ (Type ("fun", (Type ("list", _) :: _))) [t] =
3507
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
   198
      Syntax.const "length" $ t
4151
5c19cd418c33 adapted typed_print_translation;
wenzelm
parents: 4132
diff changeset
   199
  | size_tr' _ _ _ = raise Match;
3507
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
   200
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
   201
in
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
   202
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
   203
val typed_print_translation = [("size", size_tr')];
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
   204
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
   205
end;