src/HOL/Induct/SList.thy
author nipkow
Sat, 20 Jun 2009 01:53:39 +0200
changeset 31729 b9299916d618
parent 30166 f47c812de07c
child 32960 69916a850301
permissions -rw-r--r--
new lemmas and tuning
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
     1
(*  Title:      SList.thy
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
     3
    Author:     B. Wolff, University of Bremen
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     4
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
     5
Enriched theory of lists; mutual indirect recursive data-types.
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
     6
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
     7
Definition of type 'a list (strict lists) by a least fixed point
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     8
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     9
We use          list(A) == lfp(%Z. {NUMB(0)} <+> A <*> Z)
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    10
and not         list    == lfp(%Z. {NUMB(0)} <+> range(Leaf) <*> Z)
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
    11
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
    12
so that list can serve as a "functor" for defining other recursive types.
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
    13
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
    14
This enables the conservative construction of mutual recursive data-types
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
    15
such as
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
    16
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
    17
datatype 'a m = Node 'a * ('a m) list
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
    18
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
    19
Tidied by lcp.  Still needs removal of nat_rec.
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    20
*)
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    21
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
    22
header {* Extended List Theory (old) *}
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
    23
23235
eb365b39b36d authentic syntax for List.append
haftmann
parents: 23029
diff changeset
    24
theory SList
eb365b39b36d authentic syntax for List.append
haftmann
parents: 23029
diff changeset
    25
imports Sexp
eb365b39b36d authentic syntax for List.append
haftmann
parents: 23029
diff changeset
    26
begin
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
    27
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
    28
(*Hilbert_Choice is needed for the function "inv"*)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
    29
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
    30
(* *********************************************************************** *)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
    31
(*                                                                         *)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
    32
(* Building up data type                                                   *)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
    33
(*                                                                         *)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
    34
(* *********************************************************************** *)
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    35
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    36
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
    37
(* Defining the Concrete Constructors *)
19736
wenzelm
parents: 18413
diff changeset
    38
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
    39
  NIL  :: "'a item" where
20801
d3616b4abe1b proper import of Main HOL;
wenzelm
parents: 20770
diff changeset
    40
  "NIL = In0(Numb(0))"
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    41
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
    42
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
    43
  CONS :: "['a item, 'a item] => 'a item" where
20801
d3616b4abe1b proper import of Main HOL;
wenzelm
parents: 20770
diff changeset
    44
  "CONS M N = In1(Scons M N)"
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    45
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 23281
diff changeset
    46
inductive_set
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 23281
diff changeset
    47
  list :: "'a item set => 'a item set"
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 23281
diff changeset
    48
  for A :: "'a item set"
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 23281
diff changeset
    49
  where
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
    50
    NIL_I:  "NIL: list A"
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 23281
diff changeset
    51
  | CONS_I: "[| a: A;  M: list A |] ==> CONS a M : list A"
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    52
5977
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5191
diff changeset
    53
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5191
diff changeset
    54
typedef (List)
20801
d3616b4abe1b proper import of Main HOL;
wenzelm
parents: 20770
diff changeset
    55
    'a list = "list(range Leaf) :: 'a item set" 
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
    56
  by (blast intro: list.NIL_I)
5977
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5191
diff changeset
    57
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
    58
abbreviation "Case == Datatype.Case"
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
    59
abbreviation "Split == Datatype.Split"
20801
d3616b4abe1b proper import of Main HOL;
wenzelm
parents: 20770
diff changeset
    60
19736
wenzelm
parents: 18413
diff changeset
    61
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
    62
  List_case :: "['b, ['a item, 'a item]=>'b, 'a item] => 'b" where
20801
d3616b4abe1b proper import of Main HOL;
wenzelm
parents: 20770
diff changeset
    63
  "List_case c d = Case(%x. c)(Split(d))"
5977
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5191
diff changeset
    64
  
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
    65
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
    66
  List_rec  :: "['a item, 'b, ['a item, 'a item, 'b]=>'b] => 'b" where
22267
ea31e6ea0e2e Adapted to changes in Transitive_Closure theory.
berghofe
parents: 21404
diff changeset
    67
  "List_rec M c d = wfrec (pred_sexp^+)
20801
d3616b4abe1b proper import of Main HOL;
wenzelm
parents: 20770
diff changeset
    68
                           (%g. List_case c (%x y. d x y (g y))) M"
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    69
5977
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5191
diff changeset
    70
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
    71
(* *********************************************************************** *)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
    72
(*                                                                         *)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
    73
(* Abstracting data type                                                   *)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
    74
(*                                                                         *)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
    75
(* *********************************************************************** *)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
    76
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
    77
(*Declaring the abstract list constructors*)
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
    78
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
    79
no_translations
20801
d3616b4abe1b proper import of Main HOL;
wenzelm
parents: 20770
diff changeset
    80
  "[x, xs]" == "x#[xs]"
d3616b4abe1b proper import of Main HOL;
wenzelm
parents: 20770
diff changeset
    81
  "[x]" == "x#[]"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
    82
no_notation
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
    83
  Nil  ("[]") and
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
    84
  Cons (infixr "#" 65)
20801
d3616b4abe1b proper import of Main HOL;
wenzelm
parents: 20770
diff changeset
    85
19736
wenzelm
parents: 18413
diff changeset
    86
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
    87
  Nil       :: "'a list"                               ("[]") where
19736
wenzelm
parents: 18413
diff changeset
    88
   "Nil  = Abs_List(NIL)"
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
    89
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
    90
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
    91
  "Cons"       :: "['a, 'a list] => 'a list"           (infixr "#" 65) where
19736
wenzelm
parents: 18413
diff changeset
    92
   "x#xs = Abs_List(CONS (Leaf x)(Rep_List xs))"
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
    93
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
    94
definition
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
    95
  (* list Recursion -- the trancl is Essential; see list.ML *)
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
    96
  list_rec  :: "['a list, 'b, ['a, 'a list, 'b]=>'b] => 'b" where
19736
wenzelm
parents: 18413
diff changeset
    97
   "list_rec l c d =
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
    98
      List_rec(Rep_List l) c (%x y r. d(inv Leaf x)(Abs_List y) r)"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
    99
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   100
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   101
  list_case :: "['b, ['a, 'a list]=>'b, 'a list] => 'b" where
19736
wenzelm
parents: 18413
diff changeset
   102
   "list_case a f xs = list_rec xs a (%x xs r. f x xs)"
14653
0848ab6fe5fc constdefs: proper order;
wenzelm
parents: 13612
diff changeset
   103
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   104
(* list Enumeration *)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   105
translations
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   106
  "[x, xs]" == "x#[xs]"
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   107
  "[x]"     == "x#[]"
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   108
20770
2c583720436e fixed translations: CONST;
wenzelm
parents: 19736
diff changeset
   109
  "case xs of [] => a | y#ys => b" == "CONST list_case(a, %y ys. b, xs)"
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   110
20770
2c583720436e fixed translations: CONST;
wenzelm
parents: 19736
diff changeset
   111
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   112
(* *********************************************************************** *)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   113
(*                                                                         *)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   114
(* Generalized Map Functionals                                             *)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   115
(*                                                                         *)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   116
(* *********************************************************************** *)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   117
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   118
  
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   119
(* Generalized Map Functionals *)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   120
19736
wenzelm
parents: 18413
diff changeset
   121
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   122
  Rep_map   :: "('b => 'a item) => ('b list => 'a item)" where
19736
wenzelm
parents: 18413
diff changeset
   123
   "Rep_map f xs = list_rec xs  NIL(%x l r. CONS(f x) r)"
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   124
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   125
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   126
  Abs_map   :: "('a item => 'b) => 'a item => 'b list" where
19736
wenzelm
parents: 18413
diff changeset
   127
   "Abs_map g M  = List_rec M Nil (%N L r. g(N)#r)"
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   128
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   129
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   130
(**** Function definitions ****)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   131
19736
wenzelm
parents: 18413
diff changeset
   132
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   133
  null      :: "'a list => bool" where
19736
wenzelm
parents: 18413
diff changeset
   134
  "null xs  = list_rec xs True (%x xs r. False)"
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   135
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   136
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   137
  hd        :: "'a list => 'a" where
19736
wenzelm
parents: 18413
diff changeset
   138
  "hd xs    = list_rec xs (@x. True) (%x xs r. x)"
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   139
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   140
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   141
  tl        :: "'a list => 'a list" where
19736
wenzelm
parents: 18413
diff changeset
   142
  "tl xs    = list_rec xs (@xs. True) (%x xs r. xs)"
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   143
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   144
definition
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   145
  (* a total version of tl: *)
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   146
  ttl       :: "'a list => 'a list" where
19736
wenzelm
parents: 18413
diff changeset
   147
  "ttl xs   = list_rec xs [] (%x xs r. xs)"
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   148
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   149
no_notation member  (infixl "mem" 55)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   150
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   151
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   152
  member :: "['a, 'a list] => bool"    (infixl "mem" 55) where
19736
wenzelm
parents: 18413
diff changeset
   153
  "x mem xs = list_rec xs False (%y ys r. if y=x then True else r)"
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   154
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   155
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   156
  list_all  :: "('a => bool) => ('a list => bool)" where
19736
wenzelm
parents: 18413
diff changeset
   157
  "list_all P xs = list_rec xs True(%x l r. P(x) & r)"
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   158
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   159
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   160
  map       :: "('a=>'b) => ('a list => 'b list)" where
19736
wenzelm
parents: 18413
diff changeset
   161
  "map f xs = list_rec xs [] (%x l r. f(x)#r)"
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   162
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   163
no_notation append  (infixr "@" 65)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   164
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   165
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   166
  append    :: "['a list, 'a list] => 'a list"   (infixr "@" 65) where
19736
wenzelm
parents: 18413
diff changeset
   167
  "xs@ys = list_rec xs ys (%x l r. x#r)"
5977
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5191
diff changeset
   168
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   169
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   170
  filter    :: "['a => bool, 'a list] => 'a list" where
19736
wenzelm
parents: 18413
diff changeset
   171
  "filter P xs = list_rec xs []  (%x xs r. if P(x)then x#r else r)"
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   172
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   173
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   174
  foldl     :: "[['b,'a] => 'b, 'b, 'a list] => 'b" where
19736
wenzelm
parents: 18413
diff changeset
   175
  "foldl f a xs = list_rec xs (%a. a)(%x xs r.%a. r(f a x))(a)"
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   176
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   177
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   178
  foldr     :: "[['a,'b] => 'b, 'b, 'a list] => 'b" where
19736
wenzelm
parents: 18413
diff changeset
   179
  "foldr f a xs     = list_rec xs a (%x xs r. (f x r))"
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   180
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   181
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   182
  length    :: "'a list => nat" where
19736
wenzelm
parents: 18413
diff changeset
   183
  "length xs = list_rec xs 0 (%x xs r. Suc r)"
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   184
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   185
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   186
  drop      :: "['a list,nat] => 'a list" where
19736
wenzelm
parents: 18413
diff changeset
   187
  "drop t n = (nat_rec(%x. x)(%m r xs. r(ttl xs)))(n)(t)"
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   188
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   189
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   190
  copy      :: "['a, nat] => 'a list"  where     (* make list of n copies of x *)
19736
wenzelm
parents: 18413
diff changeset
   191
  "copy t   = nat_rec [] (%m xs. t # xs)"
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   192
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   193
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   194
  flat      :: "'a list list => 'a list" where
19736
wenzelm
parents: 18413
diff changeset
   195
  "flat     = foldr (op @) []"
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   196
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   197
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   198
  nth       :: "[nat, 'a list] => 'a" where
19736
wenzelm
parents: 18413
diff changeset
   199
  "nth      = nat_rec hd (%m r xs. r(tl xs))"
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   200
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   201
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   202
  rev       :: "'a list => 'a list" where
19736
wenzelm
parents: 18413
diff changeset
   203
  "rev xs   = list_rec xs [] (%x xs xsa. xsa @ [x])"
5977
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5191
diff changeset
   204
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   205
(* miscellaneous definitions *)
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   206
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   207
  zipWith   :: "['a * 'b => 'c, 'a list * 'b list] => 'c list" where
19736
wenzelm
parents: 18413
diff changeset
   208
  "zipWith f S = (list_rec (fst S)  (%T.[])
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   209
                            (%x xs r. %T. if null T then [] 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   210
                                          else f(x,hd T) # r(tl T)))(snd(S))"
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   211
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   212
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   213
  zip       :: "'a list * 'b list => ('a*'b) list" where
19736
wenzelm
parents: 18413
diff changeset
   214
  "zip      = zipWith  (%s. s)"
14653
0848ab6fe5fc constdefs: proper order;
wenzelm
parents: 13612
diff changeset
   215
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   216
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   217
  unzip     :: "('a*'b) list => ('a list * 'b list)" where
19736
wenzelm
parents: 18413
diff changeset
   218
  "unzip    = foldr(% (a,b)(c,d).(a#c,b#d))([],[])"
5977
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5191
diff changeset
   219
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5191
diff changeset
   220
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   221
consts take      :: "['a list,nat] => 'a list"
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   222
primrec
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   223
  take_0:  "take xs 0 = []"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   224
  take_Suc: "take xs (Suc n) = list_case [] (%x l. x # take l n) xs"
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   225
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   226
consts enum      :: "[nat,nat] => nat list"
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   227
primrec
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   228
  enum_0:   "enum i 0 = []"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   229
  enum_Suc: "enum i (Suc j) = (if i <= j then enum i j @ [j] else [])"
5977
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5191
diff changeset
   230
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5191
diff changeset
   231
20801
d3616b4abe1b proper import of Main HOL;
wenzelm
parents: 20770
diff changeset
   232
no_translations
23281
e26ec695c9b3 changed filter syntax from : to <-
nipkow
parents: 23235
diff changeset
   233
  "[x\<leftarrow>xs . P]" == "filter (%x. P) xs"
20801
d3616b4abe1b proper import of Main HOL;
wenzelm
parents: 20770
diff changeset
   234
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   235
syntax
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   236
  (* Special syntax for list_all and filter *)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   237
  "@Alls"       :: "[idt, 'a list, bool] => bool"        ("(2Alls _:_./ _)" 10)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11481
diff changeset
   238
5977
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5191
diff changeset
   239
translations
23281
e26ec695c9b3 changed filter syntax from : to <-
nipkow
parents: 23235
diff changeset
   240
  "[x\<leftarrow>xs. P]" == "CONST filter(%x. P) xs"
20770
2c583720436e fixed translations: CONST;
wenzelm
parents: 19736
diff changeset
   241
  "Alls x:xs. P" == "CONST list_all(%x. P)xs"
5977
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5191
diff changeset
   242
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   243
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   244
lemma ListI: "x : list (range Leaf) ==> x : List"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   245
by (simp add: List_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   246
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   247
lemma ListD: "x : List ==> x : list (range Leaf)"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   248
by (simp add: List_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   249
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   250
lemma list_unfold: "list(A) = usum {Numb(0)} (uprod A (list(A)))"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   251
  by (fast intro!: list.intros [unfolded NIL_def CONS_def]
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   252
           elim: list.cases [unfolded NIL_def CONS_def])
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   253
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   254
(*This justifies using list in other recursive type definitions*)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   255
lemma list_mono: "A<=B ==> list(A) <= list(B)"
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 23281
diff changeset
   256
apply (rule subsetI)
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 23281
diff changeset
   257
apply (erule list.induct)
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 23281
diff changeset
   258
apply (auto intro!: list.intros)
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   259
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   260
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   261
(*Type checking -- list creates well-founded sets*)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   262
lemma list_sexp: "list(sexp) <= sexp"
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 23281
diff changeset
   263
apply (rule subsetI)
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 23281
diff changeset
   264
apply (erule list.induct)
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 23281
diff changeset
   265
apply (unfold NIL_def CONS_def)
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 23281
diff changeset
   266
apply (auto intro: sexp.intros sexp_In0I sexp_In1I)
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   267
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   268
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   269
(* A <= sexp ==> list(A) <= sexp *)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   270
lemmas list_subset_sexp = subset_trans [OF list_mono list_sexp] 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   271
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   272
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   273
(*Induction for the type 'a list *)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   274
lemma list_induct:
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   275
    "[| P(Nil);    
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   276
        !!x xs. P(xs) ==> P(x # xs) |]  ==> P(l)"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   277
apply (unfold Nil_def Cons_def) 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   278
apply (rule Rep_List_inverse [THEN subst])
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   279
			 (*types force good instantiation*)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   280
apply (rule Rep_List [unfolded List_def, THEN list.induct], simp)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   281
apply (erule Abs_List_inverse [unfolded List_def, THEN subst], blast) 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   282
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   283
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   284
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   285
(*** Isomorphisms ***)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   286
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   287
lemma inj_on_Abs_list: "inj_on Abs_List (list(range Leaf))"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   288
apply (rule inj_on_inverseI)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   289
apply (erule Abs_List_inverse [unfolded List_def])
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   290
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   291
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   292
(** Distinctness of constructors **)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   293
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   294
lemma CONS_not_NIL [iff]: "CONS M N ~= NIL"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   295
by (simp add: NIL_def CONS_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   296
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   297
lemmas NIL_not_CONS [iff] = CONS_not_NIL [THEN not_sym]
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   298
lemmas CONS_neq_NIL = CONS_not_NIL [THEN notE, standard]
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   299
lemmas NIL_neq_CONS = sym [THEN CONS_neq_NIL]
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   300
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   301
lemma Cons_not_Nil [iff]: "x # xs ~= Nil"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   302
apply (unfold Nil_def Cons_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   303
apply (rule CONS_not_NIL [THEN inj_on_Abs_list [THEN inj_on_contraD]])
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   304
apply (simp_all add: list.intros rangeI Rep_List [unfolded List_def])
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   305
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   306
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   307
lemmas Nil_not_Cons [iff] = Cons_not_Nil [THEN not_sym, standard]
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   308
lemmas Cons_neq_Nil = Cons_not_Nil [THEN notE, standard]
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   309
lemmas Nil_neq_Cons = sym [THEN Cons_neq_Nil]
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   310
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   311
(** Injectiveness of CONS and Cons **)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   312
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   313
lemma CONS_CONS_eq [iff]: "(CONS K M)=(CONS L N) = (K=L & M=N)"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   314
by (simp add: CONS_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   315
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   316
(*For reasoning about abstract list constructors*)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   317
declare Rep_List [THEN ListD, intro] ListI [intro]
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   318
declare list.intros [intro,simp]
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   319
declare Leaf_inject [dest!]
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   320
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   321
lemma Cons_Cons_eq [iff]: "(x#xs=y#ys) = (x=y & xs=ys)"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   322
apply (simp add: Cons_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   323
apply (subst Abs_List_inject)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   324
apply (auto simp add: Rep_List_inject)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   325
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   326
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   327
lemmas Cons_inject2 = Cons_Cons_eq [THEN iffD1, THEN conjE, standard]
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   328
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   329
lemma CONS_D: "CONS M N: list(A) ==> M: A & N: list(A)"
18413
50c0c118e96d removed obsolete/unused setup_induction;
wenzelm
parents: 16417
diff changeset
   330
  by (induct L == "CONS M N" set: list) auto
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   331
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   332
lemma sexp_CONS_D: "CONS M N: sexp ==> M: sexp & N: sexp"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   333
apply (simp add: CONS_def In1_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   334
apply (fast dest!: Scons_D)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   335
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   336
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   337
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   338
(*Reasoning about constructors and their freeness*)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   339
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   340
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   341
lemma not_CONS_self: "N: list(A) ==> !M. N ~= CONS M N"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   342
apply (erule list.induct) apply simp_all done
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   343
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   344
lemma not_Cons_self2: "\<forall>x. l ~= x#l"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   345
by (induct l rule: list_induct) simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   346
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   347
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   348
lemma neq_Nil_conv2: "(xs ~= []) = (\<exists>y ys. xs = y#ys)"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   349
by (induct xs rule: list_induct) auto
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   350
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   351
(** Conversion rules for List_case: case analysis operator **)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   352
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   353
lemma List_case_NIL [simp]: "List_case c h NIL = c"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   354
by (simp add: List_case_def NIL_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   355
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   356
lemma List_case_CONS [simp]: "List_case c h (CONS M N) = h M N"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   357
by (simp add: List_case_def CONS_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   358
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   359
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   360
(*** List_rec -- by wf recursion on pred_sexp ***)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   361
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   362
(* The trancl(pred_sexp) is essential because pred_sexp_CONS_I1,2 would not
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   363
   hold if pred_sexp^+ were changed to pred_sexp. *)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   364
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   365
lemma List_rec_unfold_lemma:
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   366
     "(%M. List_rec M c d) == 
22267
ea31e6ea0e2e Adapted to changes in Transitive_Closure theory.
berghofe
parents: 21404
diff changeset
   367
      wfrec (pred_sexp^+) (%g. List_case c (%x y. d x y (g y)))"
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   368
by (simp add: List_rec_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   369
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   370
lemmas List_rec_unfold = 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   371
    def_wfrec [OF List_rec_unfold_lemma wf_pred_sexp [THEN wf_trancl], 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   372
               standard]
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   373
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   374
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   375
(** pred_sexp lemmas **)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   376
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   377
lemma pred_sexp_CONS_I1: 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   378
    "[| M: sexp;  N: sexp |] ==> (M, CONS M N) : pred_sexp^+"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   379
by (simp add: CONS_def In1_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   380
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   381
lemma pred_sexp_CONS_I2: 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   382
    "[| M: sexp;  N: sexp |] ==> (N, CONS M N) : pred_sexp^+"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   383
by (simp add: CONS_def In1_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   384
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   385
lemma pred_sexp_CONS_D: 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   386
    "(CONS M1 M2, N) : pred_sexp^+ ==>  
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   387
     (M1,N) : pred_sexp^+ & (M2,N) : pred_sexp^+"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   388
apply (frule pred_sexp_subset_Sigma [THEN trancl_subset_Sigma, THEN subsetD])
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   389
apply (blast dest!: sexp_CONS_D intro: pred_sexp_CONS_I1 pred_sexp_CONS_I2 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   390
                    trans_trancl [THEN transD])
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   391
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   392
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   393
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   394
(** Conversion rules for List_rec **)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   395
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   396
lemma List_rec_NIL [simp]: "List_rec NIL c h = c"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   397
apply (rule List_rec_unfold [THEN trans])
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   398
apply (simp add: List_case_NIL)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   399
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   400
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   401
lemma List_rec_CONS [simp]:
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   402
     "[| M: sexp;  N: sexp |] 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   403
      ==> List_rec (CONS M N) c h = h M N (List_rec N c h)"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   404
apply (rule List_rec_unfold [THEN trans])
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   405
apply (simp add: pred_sexp_CONS_I2)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   406
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   407
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   408
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   409
(*** list_rec -- by List_rec ***)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   410
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   411
lemmas Rep_List_in_sexp =
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   412
    subsetD [OF range_Leaf_subset_sexp [THEN list_subset_sexp]
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   413
                Rep_List [THEN ListD]] 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   414
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   415
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   416
lemma list_rec_Nil [simp]: "list_rec Nil c h = c"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   417
by (simp add: list_rec_def ListI [THEN Abs_List_inverse] Nil_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   418
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   419
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   420
lemma list_rec_Cons [simp]: "list_rec (a#l) c h = h a l (list_rec l c h)"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   421
by (simp add: list_rec_def ListI [THEN Abs_List_inverse] Cons_def
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   422
              Rep_List_inverse Rep_List [THEN ListD] inj_Leaf Rep_List_in_sexp)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   423
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   424
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   425
(*Type checking.  Useful?*)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   426
lemma List_rec_type:
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   427
     "[| M: list(A);      
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   428
         A<=sexp;         
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   429
         c: C(NIL);       
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   430
         !!x y r. [| x: A;  y: list(A);  r: C(y) |] ==> h x y r: C(CONS x y)  
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   431
      |] ==> List_rec M c h : C(M :: 'a item)"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   432
apply (erule list.induct, simp) 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   433
apply (insert list_subset_sexp) 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   434
apply (subst List_rec_CONS, blast+)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   435
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   436
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   437
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   438
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   439
(** Generalized map functionals **)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   440
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   441
lemma Rep_map_Nil [simp]: "Rep_map f Nil = NIL"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   442
by (simp add: Rep_map_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   443
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   444
lemma Rep_map_Cons [simp]: 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   445
    "Rep_map f(x#xs) = CONS(f x)(Rep_map f xs)"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   446
by (simp add: Rep_map_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   447
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   448
lemma Rep_map_type: "(!!x. f(x): A) ==> Rep_map f xs: list(A)"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   449
apply (simp add: Rep_map_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   450
apply (rule list_induct, auto)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   451
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   452
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   453
lemma Abs_map_NIL [simp]: "Abs_map g NIL = Nil"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   454
by (simp add: Abs_map_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   455
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   456
lemma Abs_map_CONS [simp]: 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   457
    "[| M: sexp;  N: sexp |] ==> Abs_map g (CONS M N) = g(M) # Abs_map g N"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   458
by (simp add: Abs_map_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   459
19736
wenzelm
parents: 18413
diff changeset
   460
(*Eases the use of primitive recursion.*)
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   461
lemma def_list_rec_NilCons:
19736
wenzelm
parents: 18413
diff changeset
   462
    "[| !!xs. f(xs) = list_rec xs c h  |] 
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   463
     ==> f [] = c & f(x#xs) = h x xs (f xs)"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   464
by simp 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   465
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   466
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   467
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   468
lemma Abs_map_inverse:
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   469
     "[| M: list(A);  A<=sexp;  !!z. z: A ==> f(g(z)) = z |]  
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   470
      ==> Rep_map f (Abs_map g M) = M"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   471
apply (erule list.induct, simp_all)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   472
apply (insert list_subset_sexp) 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   473
apply (subst Abs_map_CONS, blast)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   474
apply blast 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   475
apply simp 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   476
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   477
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   478
(*Rep_map_inverse is obtained via Abs_Rep_map and map_ident*)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   479
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   480
(** list_case **)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   481
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   482
(* setting up rewrite sets *)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   483
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   484
text{*Better to have a single theorem with a conjunctive conclusion.*}
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   485
declare def_list_rec_NilCons [OF list_case_def, simp]
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   486
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   487
(** list_case **)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   488
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   489
lemma expand_list_case: 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   490
 "P(list_case a f xs) = ((xs=[] --> P a ) & (!y ys. xs=y#ys --> P(f y ys)))"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   491
by (induct xs rule: list_induct) simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   492
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   493
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   494
(**** Function definitions ****)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   495
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   496
declare def_list_rec_NilCons [OF null_def, simp]
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   497
declare def_list_rec_NilCons [OF hd_def, simp]
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   498
declare def_list_rec_NilCons [OF tl_def, simp]
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   499
declare def_list_rec_NilCons [OF ttl_def, simp]
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   500
declare def_list_rec_NilCons [OF append_def, simp]
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   501
declare def_list_rec_NilCons [OF member_def, simp]
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   502
declare def_list_rec_NilCons [OF map_def, simp]
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   503
declare def_list_rec_NilCons [OF filter_def, simp]
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   504
declare def_list_rec_NilCons [OF list_all_def, simp]
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   505
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   506
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   507
(** nth **)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   508
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   509
lemma def_nat_rec_0_eta:
19736
wenzelm
parents: 18413
diff changeset
   510
    "[| !!n. f = nat_rec c h |] ==> f(0) = c"
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   511
by simp
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   512
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   513
lemma def_nat_rec_Suc_eta:
19736
wenzelm
parents: 18413
diff changeset
   514
    "[| !!n. f = nat_rec c h |] ==> f(Suc(n)) = h n (f n)"
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   515
by simp
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   516
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   517
declare def_nat_rec_0_eta [OF nth_def, simp]
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   518
declare def_nat_rec_Suc_eta [OF nth_def, simp]
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   519
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   520
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   521
(** length **)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   522
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   523
lemma length_Nil [simp]: "length([]) = 0"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   524
by (simp add: length_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   525
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   526
lemma length_Cons [simp]: "length(a#xs) = Suc(length(xs))"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   527
by (simp add: length_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   528
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   529
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   530
(** @ - append **)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   531
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   532
lemma append_assoc [simp]: "(xs@ys)@zs = xs@(ys@zs)"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   533
by (induct xs rule: list_induct) simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   534
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   535
lemma append_Nil2 [simp]: "xs @ [] = xs"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   536
by (induct xs rule: list_induct) simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   537
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   538
(** mem **)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   539
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   540
lemma mem_append [simp]: "x mem (xs@ys) = (x mem xs | x mem ys)"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   541
by (induct xs rule: list_induct) simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   542
23281
e26ec695c9b3 changed filter syntax from : to <-
nipkow
parents: 23235
diff changeset
   543
lemma mem_filter [simp]: "x mem [x\<leftarrow>xs. P x ] = (x mem xs & P(x))"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   544
by (induct xs rule: list_induct) simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   545
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   546
(** list_all **)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   547
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   548
lemma list_all_True [simp]: "(Alls x:xs. True) = True"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   549
by (induct xs rule: list_induct) simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   550
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   551
lemma list_all_conj [simp]:
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   552
     "list_all p (xs@ys) = ((list_all p xs) & (list_all p ys))"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   553
by (induct xs rule: list_induct) simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   554
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   555
lemma list_all_mem_conv: "(Alls x:xs. P(x)) = (!x. x mem xs --> P(x))"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   556
apply (induct xs rule: list_induct)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   557
apply simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   558
apply blast 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   559
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   560
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   561
lemma nat_case_dist : "(! n. P n) = (P 0 & (! n. P (Suc n)))"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   562
apply auto
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   563
apply (induct_tac n)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   564
apply auto
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   565
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   566
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   567
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   568
lemma alls_P_eq_P_nth: "(Alls u:A. P u) = (!n. n < length A --> P(nth n A))"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   569
apply (induct_tac A rule: list_induct)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   570
apply simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   571
apply (rule trans)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   572
apply (rule_tac [2] nat_case_dist [symmetric], simp_all)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   573
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   574
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   575
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   576
lemma list_all_imp:
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   577
     "[| !x. P x --> Q x;  (Alls x:xs. P(x)) |] ==> (Alls x:xs. Q(x))"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   578
by (simp add: list_all_mem_conv)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   579
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   580
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   581
(** The functional "map" and the generalized functionals **)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   582
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   583
lemma Abs_Rep_map: 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   584
     "(!!x. f(x): sexp) ==>  
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   585
        Abs_map g (Rep_map f xs) = map (%t. g(f(t))) xs"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   586
apply (induct xs rule: list_induct)
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   587
apply (simp_all add: Rep_map_type list_sexp [THEN subsetD])
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   588
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   589
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   590
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   591
(** Additional mapping lemmas **)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   592
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   593
lemma map_ident [simp]: "map(%x. x)(xs) = xs"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   594
by (induct xs rule: list_induct) simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   595
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   596
lemma map_append [simp]: "map f (xs@ys) = map f xs  @ map f ys"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   597
by (induct xs rule: list_induct) simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   598
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   599
lemma map_compose: "map(f o g)(xs) = map f (map g xs)"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   600
apply (simp add: o_def)
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   601
apply (induct xs rule: list_induct)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   602
apply simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   603
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   604
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   605
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   606
lemma mem_map_aux1 [rule_format]:
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   607
     "x mem (map f q) --> (\<exists>y. y mem q & x = f y)"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   608
by (induct q rule: list_induct) auto
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   609
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   610
lemma mem_map_aux2 [rule_format]: 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   611
     "(\<exists>y. y mem q & x = f y) --> x mem (map f q)"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   612
by (induct q rule: list_induct) auto
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   613
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   614
lemma mem_map: "x mem (map f q) = (\<exists>y. y mem q & x = f y)"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   615
apply (rule iffI)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   616
apply (erule mem_map_aux1)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   617
apply (erule mem_map_aux2)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   618
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   619
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   620
lemma hd_append [rule_format]: "A ~= [] --> hd(A @ B) = hd(A)"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   621
by (induct A rule: list_induct) auto
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   622
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   623
lemma tl_append [rule_format]: "A ~= [] --> tl(A @ B) = tl(A) @ B"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   624
by (induct A rule: list_induct) auto
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   625
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   626
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   627
(** take **)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   628
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   629
lemma take_Suc1 [simp]: "take [] (Suc x) = []"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   630
by simp
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   631
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   632
lemma take_Suc2 [simp]: "take(a#xs)(Suc x) = a#take xs x"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   633
by simp
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   634
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   635
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   636
(** drop **)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   637
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   638
lemma drop_0 [simp]: "drop xs 0 = xs"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   639
by (simp add: drop_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   640
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   641
lemma drop_Suc1 [simp]: "drop [] (Suc x) = []"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   642
apply (induct x) 
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   643
apply (simp_all add: drop_def)
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   644
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   645
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   646
lemma drop_Suc2 [simp]: "drop(a#xs)(Suc x) = drop xs x"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   647
by (simp add: drop_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   648
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   649
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   650
(** copy **)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   651
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   652
lemma copy_0 [simp]: "copy x 0 = []"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   653
by (simp add: copy_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   654
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   655
lemma copy_Suc [simp]: "copy x (Suc y) = x # copy x y"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   656
by (simp add: copy_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   657
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   658
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   659
(** fold **)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   660
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   661
lemma foldl_Nil [simp]: "foldl f a [] = a"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   662
by (simp add: foldl_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   663
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   664
lemma foldl_Cons [simp]: "foldl f a(x#xs) = foldl f (f a x) xs"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   665
by (simp add: foldl_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   666
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   667
lemma foldr_Nil [simp]: "foldr f a [] = a"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   668
by (simp add: foldr_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   669
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   670
lemma foldr_Cons [simp]: "foldr f z(x#xs)  = f x (foldr f z xs)"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   671
by (simp add: foldr_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   672
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   673
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   674
(** flat **)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   675
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   676
lemma flat_Nil [simp]: "flat [] = []"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   677
by (simp add: flat_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   678
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   679
lemma flat_Cons [simp]: "flat (x # xs) = x @ flat xs"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   680
by (simp add: flat_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   681
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   682
(** rev **)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   683
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   684
lemma rev_Nil [simp]: "rev [] = []"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   685
by (simp add: rev_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   686
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   687
lemma rev_Cons [simp]: "rev (x # xs) = rev xs @ [x]"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   688
by (simp add: rev_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   689
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   690
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   691
(** zip **)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   692
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   693
lemma zipWith_Cons_Cons [simp]:
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   694
     "zipWith f (a#as,b#bs)   = f(a,b) # zipWith f (as,bs)"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   695
by (simp add: zipWith_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   696
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   697
lemma zipWith_Nil_Nil [simp]: "zipWith f ([],[])      = []"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   698
by (simp add: zipWith_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   699
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   700
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   701
lemma zipWith_Cons_Nil [simp]: "zipWith f (x,[])  = []"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   702
by (induct x rule: list_induct) (simp_all add: zipWith_def)
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   703
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   704
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   705
lemma zipWith_Nil_Cons [simp]: "zipWith f ([],x) = []"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   706
by (simp add: zipWith_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   707
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   708
lemma unzip_Nil [simp]: "unzip [] = ([],[])"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   709
by (simp add: unzip_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   710
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   711
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   712
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   713
(** SOME LIST THEOREMS **)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   714
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   715
(* SQUIGGOL LEMMAS *)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   716
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   717
lemma map_compose_ext: "map(f o g) = ((map f) o (map g))"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   718
apply (simp add: o_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   719
apply (rule ext)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   720
apply (simp add: map_compose [symmetric] o_def)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   721
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   722
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   723
lemma map_flat: "map f (flat S) = flat(map (map f) S)"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   724
by (induct S rule: list_induct) simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   725
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   726
lemma list_all_map_eq: "(Alls u:xs. f(u) = g(u)) --> map f xs = map g xs"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   727
by (induct xs rule: list_induct) simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   728
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   729
lemma filter_map_d: "filter p (map f xs) = map f (filter(p o f)(xs))"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   730
by (induct xs rule: list_induct) simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   731
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   732
lemma filter_compose: "filter p (filter q xs) = filter(%x. p x & q x) xs"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   733
by (induct xs rule: list_induct) simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   734
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   735
(* "filter(p, filter(q,xs)) = filter(q, filter(p,xs))",
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   736
   "filter(p, filter(p,xs)) = filter(p,xs)" BIRD's thms.*)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   737
 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   738
lemma filter_append [rule_format, simp]:
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   739
     "\<forall>B. filter p (A @ B) = (filter p A @ filter p B)"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   740
by (induct A rule: list_induct) simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   741
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   742
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   743
(* inits(xs) == map(fst,splits(xs)), 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   744
 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   745
   splits([]) = []
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   746
   splits(a # xs) = <[],xs> @ map(%x. <a # fst(x),snd(x)>, splits(xs))
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   747
   (x @ y = z) = <x,y> mem splits(z)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   748
   x mem xs & y mem ys = <x,y> mem diag(xs,ys) *)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   749
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   750
lemma length_append: "length(xs@ys) = length(xs)+length(ys)"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   751
by (induct xs rule: list_induct) simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   752
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   753
lemma length_map: "length(map f xs) = length(xs)"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   754
by (induct xs rule: list_induct) simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   755
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   756
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   757
lemma take_Nil [simp]: "take [] n = []"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   758
by (induct n) simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   759
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   760
lemma take_take_eq [simp]: "\<forall>n. take (take xs n) n = take xs n"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   761
apply (induct xs rule: list_induct)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   762
apply simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   763
apply (rule allI)
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   764
apply (induct_tac n)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   765
apply auto
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   766
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   767
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   768
lemma take_take_Suc_eq1 [rule_format]:
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   769
     "\<forall>n. take (take xs(Suc(n+m))) n = take xs n"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   770
apply (induct_tac xs rule: list_induct)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   771
apply simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   772
apply (rule allI)
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   773
apply (induct_tac n)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   774
apply auto
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   775
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   776
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   777
declare take_Suc [simp del]
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   778
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   779
lemma take_take_1: "take (take xs (n+m)) n = take xs n"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   780
apply (induct m)
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   781
apply (simp_all add: take_take_Suc_eq1)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   782
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   783
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   784
lemma take_take_Suc_eq2 [rule_format]:
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   785
     "\<forall>n. take (take xs n)(Suc(n+m)) = take xs n"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   786
apply (induct_tac xs rule: list_induct)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   787
apply simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   788
apply (rule allI)
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   789
apply (induct_tac n)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   790
apply auto
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   791
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   792
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   793
lemma take_take_2: "take(take xs n)(n+m) = take xs n"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   794
apply (induct m)
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   795
apply (simp_all add: take_take_Suc_eq2)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   796
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   797
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   798
(* length(take(xs,n)) = min(n, length(xs)) *)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   799
(* length(drop(xs,n)) = length(xs) - n *)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   800
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   801
lemma drop_Nil [simp]: "drop  [] n  = []"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   802
by (induct n) auto
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   803
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   804
lemma drop_drop [rule_format]: "\<forall>xs. drop (drop xs m) n = drop xs(m+n)"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   805
apply (induct_tac m)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   806
apply auto
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   807
apply (induct_tac xs rule: list_induct)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   808
apply auto
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   809
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   810
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   811
lemma take_drop [rule_format]: "\<forall>xs. (take xs n) @ (drop xs n) = xs"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   812
apply (induct_tac n)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   813
apply auto
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   814
apply (induct_tac xs rule: list_induct)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   815
apply auto
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   816
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   817
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   818
lemma copy_copy: "copy x n @ copy x m = copy x (n+m)"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   819
by (induct n) auto
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   820
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   821
lemma length_copy: "length(copy x n)  = n"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   822
by (induct n) auto
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   823
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   824
lemma length_take [rule_format, simp]:
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   825
     "\<forall>xs. length(take xs n) = min (length xs) n"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   826
apply (induct n)
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   827
 apply auto
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   828
apply (induct_tac xs rule: list_induct)
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   829
 apply auto
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   830
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   831
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   832
lemma length_take_drop: "length(take A k) + length(drop A k) = length(A)" 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   833
by (simp only: length_append [symmetric] take_drop)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   834
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   835
lemma take_append [rule_format]: "\<forall>A. length(A) = n --> take(A@B) n = A"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   836
apply (induct n)
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   837
apply (rule allI)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   838
apply (rule_tac [2] allI)
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   839
apply (induct_tac A rule: list_induct)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   840
apply (induct_tac [3] A rule: list_induct, simp_all)
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   841
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   842
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   843
lemma take_append2 [rule_format]:
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   844
     "\<forall>A. length(A) = n --> take(A@B) (n+k) = A @ take B k"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   845
apply (induct n)
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   846
apply (rule allI)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   847
apply (rule_tac [2] allI)
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   848
apply (induct_tac A rule: list_induct)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   849
apply (induct_tac [3] A rule: list_induct, simp_all)
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   850
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   851
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   852
lemma take_map [rule_format]: "\<forall>n. take (map f A) n = map f (take A n)"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   853
apply (induct A rule: list_induct)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   854
apply simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   855
apply (rule allI)
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   856
apply (induct_tac n)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   857
apply simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   858
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   859
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   860
lemma drop_append [rule_format]: "\<forall>A. length(A) = n --> drop(A@B)n = B"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   861
apply (induct n)
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   862
apply (rule allI)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   863
apply (rule_tac [2] allI)
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   864
apply (induct_tac A rule: list_induct)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   865
apply (induct_tac [3] A rule: list_induct)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   866
apply simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   867
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   868
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   869
lemma drop_append2 [rule_format]:
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   870
     "\<forall>A. length(A) = n --> drop(A@B)(n+k) = drop B k"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   871
apply (induct n)
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   872
apply (rule allI)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   873
apply (rule_tac [2] allI)
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   874
apply (induct_tac A rule: list_induct)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   875
apply (induct_tac [3] A rule: list_induct)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   876
apply simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   877
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   878
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   879
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   880
lemma drop_all [rule_format]: "\<forall>A. length(A) = n --> drop A n = []"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   881
apply (induct n)
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   882
apply (rule allI)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   883
apply (rule_tac [2] allI)
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   884
apply (induct_tac A rule: list_induct)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   885
apply (induct_tac [3] A rule: list_induct)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   886
apply auto
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   887
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   888
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   889
lemma drop_map [rule_format]: "\<forall>n. drop (map f A) n = map f (drop A n)"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   890
apply (induct A rule: list_induct)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   891
apply simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   892
apply (rule allI)
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   893
apply (induct_tac n)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   894
apply simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   895
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   896
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   897
lemma take_all [rule_format]: "\<forall>A. length(A) = n --> take A n = A"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   898
apply (induct n)
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   899
apply (rule allI)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   900
apply (rule_tac [2] allI)
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   901
apply (induct_tac A rule: list_induct)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   902
apply (induct_tac [3] A rule: list_induct)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   903
apply auto
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   904
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   905
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   906
lemma foldl_single: "foldl f a [b] = f a b"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   907
by simp_all
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   908
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   909
lemma foldl_append [simp]:
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   910
  "\<And>a. foldl f a (A @ B) = foldl f (foldl f a A) B"
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   911
by (induct A rule: list_induct) simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   912
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   913
lemma foldl_map:
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   914
  "\<And>e. foldl f e (map g S) = foldl (%x y. f x (g y)) e S"
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   915
by (induct S rule: list_induct) simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   916
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   917
lemma foldl_neutr_distr [rule_format]:
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   918
  assumes r_neutr: "\<forall>a. f a e = a" 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   919
      and r_neutl: "\<forall>a. f e a = a"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   920
      and assoc:   "\<forall>a b c. f a (f b c) = f(f a b) c"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   921
  shows "\<forall>y. f y (foldl f e A) = foldl f y A"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   922
apply (induct A rule: list_induct)
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   923
apply (simp_all add: r_neutr r_neutl, clarify) 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   924
apply (erule all_dupE) 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   925
apply (rule trans) 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   926
prefer 2 apply assumption
13612
55d32e76ef4e Adapted to new simplifier.
berghofe
parents: 13079
diff changeset
   927
apply (simp (no_asm_use) add: assoc [THEN spec, THEN spec, THEN spec, THEN sym])
55d32e76ef4e Adapted to new simplifier.
berghofe
parents: 13079
diff changeset
   928
apply simp
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   929
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   930
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   931
lemma foldl_append_sym: 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   932
"[| !a. f a e = a; !a. f e a = a;           
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   933
    !a b c. f a (f b c) = f(f a b) c |]    
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   934
  ==> foldl f e (A @ B) = f(foldl f e A)(foldl f e B)"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   935
apply (rule trans)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   936
apply (rule foldl_append)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   937
apply (rule sym) 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   938
apply (rule foldl_neutr_distr, auto)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   939
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   940
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   941
lemma foldr_append [rule_format, simp]:
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   942
     "\<forall>a. foldr f a (A @ B) = foldr f (foldr f a B) A"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   943
by (induct A rule: list_induct) simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   944
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   945
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   946
lemma foldr_map: "\<And>e. foldr f e (map g S) = foldr (f o g) e S"
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   947
by (induct S rule: list_induct) (simp_all add: o_def)
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   948
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   949
lemma foldr_Un_eq_UN: "foldr op Un {} S = (UN X: {t. t mem S}.X)"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   950
by (induct S rule: list_induct) auto
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   951
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   952
lemma foldr_neutr_distr:
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   953
     "[| !a. f e a = a; !a b c. f a (f b c) = f(f a b) c |]    
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   954
      ==> foldr f y S = f (foldr f e S) y"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   955
by (induct S rule: list_induct) auto
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   956
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   957
lemma foldr_append2: 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   958
    "[| !a. f e a = a; !a b c. f a (f b c) = f(f a b) c |]
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   959
     ==> foldr f e (A @ B) = f (foldr f e A) (foldr f e B)"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   960
apply auto
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   961
apply (rule foldr_neutr_distr)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   962
apply auto
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   963
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   964
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   965
lemma foldr_flat: 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   966
    "[| !a. f e a = a; !a b c. f a (f b c) = f(f a b) c |] ==>  
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   967
      foldr f e (flat S) = (foldr f e)(map (foldr f e) S)"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   968
apply (induct S rule: list_induct)
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   969
apply (simp_all del: foldr_append add: foldr_append2)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   970
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   971
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   972
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   973
lemma list_all_map: "(Alls x:map f xs .P(x)) = (Alls x:xs.(P o f)(x))"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   974
by (induct xs rule: list_induct) auto
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   975
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   976
lemma list_all_and: 
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   977
     "(Alls x:xs. P(x)&Q(x)) = ((Alls x:xs. P(x))&(Alls x:xs. Q(x)))"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   978
by (induct xs rule: list_induct) auto
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   979
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   980
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   981
lemma nth_map [rule_format]:
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   982
     "\<forall>i. i < length(A)  --> nth i (map f A) = f(nth i A)"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   983
apply (induct A rule: list_induct)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   984
apply simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   985
apply (rule allI)
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   986
apply (induct_tac i)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   987
apply auto
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   988
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   989
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   990
lemma nth_app_cancel_right [rule_format]:
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   991
     "\<forall>i. i < length(A)  --> nth i(A@B) = nth i A"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   992
apply (induct A rule: list_induct)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   993
apply simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   994
apply (rule allI)
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   995
apply (induct_tac i)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
   996
apply simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   997
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   998
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
   999
lemma nth_app_cancel_left [rule_format]:
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
  1000
     "\<forall>n. n = length(A) --> nth(n+i)(A@B) = nth i B"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
  1001
by (induct A rule: list_induct) simp_all
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
  1002
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
  1003
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
  1004
(** flat **)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
  1005
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
  1006
lemma flat_append [simp]: "flat(xs@ys) = flat(xs) @ flat(ys)"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
  1007
by (induct xs rule: list_induct) auto
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
  1008
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
  1009
lemma filter_flat: "filter p (flat S) = flat(map (filter p) S)"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
  1010
by (induct S rule: list_induct) auto
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
  1011
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
  1012
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
  1013
(** rev **)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
  1014
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
  1015
lemma rev_append [simp]: "rev(xs@ys) = rev(ys) @ rev(xs)"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
  1016
by (induct xs rule: list_induct) auto
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
  1017
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
  1018
lemma rev_rev_ident [simp]: "rev(rev l) = l"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
  1019
by (induct l rule: list_induct) auto
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
  1020
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
  1021
lemma rev_flat: "rev(flat ls) = flat (map rev (rev ls))"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
  1022
by (induct ls rule: list_induct) auto
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
  1023
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
  1024
lemma rev_map_distrib: "rev(map f l) = map f (rev l)"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
  1025
by (induct l rule: list_induct) auto
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
  1026
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
  1027
lemma foldl_rev: "foldl f b (rev l) = foldr (%x y. f y x) b l"
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
  1028
by (induct l rule: list_induct) auto
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
  1029
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
  1030
lemma foldr_rev: "foldr f b (rev l) = foldl (%x y. f y x) b l"
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
  1031
apply (rule sym)
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
  1032
apply (rule trans)
30166
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
  1033
apply (rule_tac [2] foldl_rev)
f47c812de07c replaced low-level 'no_syntax' by 'no_notation';
wenzelm
parents: 23746
diff changeset
  1034
apply simp
13079
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
  1035
done
e7738aa7267f conversion of Induct/{Slist,Sexp} to Isar scripts
paulson
parents: 12169
diff changeset
  1036
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
  1037
end