| 
35247
 | 
     1  | 
(*  Title:      HOL/Induct/SList.thy
  | 
| 
 | 
     2  | 
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
  | 
| 
 | 
     3  | 
  | 
| 
 | 
     4  | 
This theory is strictly of historical interest. It illustrates how
  | 
| 
 | 
     5  | 
recursive datatypes can be constructed in higher-order logic from
  | 
| 
 | 
     6  | 
first principles. Isabelle's datatype package automates a
  | 
| 
 | 
     7  | 
construction of this sort.
  | 
| 
 | 
     8  | 
  | 
| 
 | 
     9  | 
Enriched theory of lists; mutual indirect recursive data-types.
  | 
| 
 | 
    10  | 
  | 
| 
 | 
    11  | 
Definition of type 'a list (strict lists) by a least fixed point
  | 
| 
 | 
    12  | 
  | 
| 
 | 
    13  | 
We use          list(A) == lfp(%Z. {NUMB(0)} <+> A <*> Z)
 | 
| 
 | 
    14  | 
and not         list    == lfp(%Z. {NUMB(0)} <+> range(Leaf) <*> Z)
 | 
| 
 | 
    15  | 
  | 
| 
 | 
    16  | 
so that list can serve as a "functor" for defining other recursive types.
  | 
| 
 | 
    17  | 
  | 
| 
 | 
    18  | 
This enables the conservative construction of mutual recursive data-types
  | 
| 
 | 
    19  | 
such as
  | 
| 
 | 
    20  | 
  | 
| 
 | 
    21  | 
datatype 'a m = Node 'a * ('a m) list
 | 
| 
 | 
    22  | 
*)
  | 
| 
 | 
    23  | 
  | 
| 
 | 
    24  | 
header {* Extended List Theory (old) *}
 | 
| 
 | 
    25  | 
  | 
| 
 | 
    26  | 
theory SList
  | 
| 
 | 
    27  | 
imports Sexp
  | 
| 
 | 
    28  | 
begin
  | 
| 
 | 
    29  | 
  | 
| 
 | 
    30  | 
(*Hilbert_Choice is needed for the function "inv"*)
  | 
| 
 | 
    31  | 
  | 
| 
 | 
    32  | 
(* *********************************************************************** *)
  | 
| 
 | 
    33  | 
(*                                                                         *)
  | 
| 
 | 
    34  | 
(* Building up data type                                                   *)
  | 
| 
 | 
    35  | 
(*                                                                         *)
  | 
| 
 | 
    36  | 
(* *********************************************************************** *)
  | 
| 
 | 
    37  | 
  | 
| 
 | 
    38  | 
  | 
| 
 | 
    39  | 
(* Defining the Concrete Constructors *)
  | 
| 
 | 
    40  | 
definition
  | 
| 
 | 
    41  | 
  NIL  :: "'a item" where
  | 
| 
 | 
    42  | 
  "NIL = In0(Numb(0))"
  | 
| 
 | 
    43  | 
  | 
| 
 | 
    44  | 
definition
  | 
| 
 | 
    45  | 
  CONS :: "['a item, 'a item] => 'a item" where
  | 
| 
 | 
    46  | 
  "CONS M N = In1(Scons M N)"
  | 
| 
 | 
    47  | 
  | 
| 
 | 
    48  | 
inductive_set
  | 
| 
 | 
    49  | 
  list :: "'a item set => 'a item set"
  | 
| 
 | 
    50  | 
  for A :: "'a item set"
  | 
| 
 | 
    51  | 
  where
  | 
| 
 | 
    52  | 
    NIL_I:  "NIL: list A"
  | 
| 
 | 
    53  | 
  | CONS_I: "[| a: A;  M: list A |] ==> CONS a M : list A"
  | 
| 
 | 
    54  | 
  | 
| 
 | 
    55  | 
  | 
| 
 | 
    56  | 
typedef (List)
  | 
| 
 | 
    57  | 
    'a list = "list(range Leaf) :: 'a item set" 
  | 
| 
 | 
    58  | 
  by (blast intro: list.NIL_I)
  | 
| 
 | 
    59  | 
  | 
| 
 | 
    60  | 
abbreviation "Case == Datatype.Case"
  | 
| 
 | 
    61  | 
abbreviation "Split == Datatype.Split"
  | 
| 
 | 
    62  | 
  | 
| 
 | 
    63  | 
definition
  | 
| 
 | 
    64  | 
  List_case :: "['b, ['a item, 'a item]=>'b, 'a item] => 'b" where
  | 
| 
 | 
    65  | 
  "List_case c d = Case(%x. c)(Split(d))"
  | 
| 
 | 
    66  | 
  
  | 
| 
 | 
    67  | 
definition
  | 
| 
 | 
    68  | 
  List_rec  :: "['a item, 'b, ['a item, 'a item, 'b]=>'b] => 'b" where
  | 
| 
 | 
    69  | 
  "List_rec M c d = wfrec (pred_sexp^+)
  | 
| 
 | 
    70  | 
                           (%g. List_case c (%x y. d x y (g y))) M"
  | 
| 
 | 
    71  | 
  | 
| 
 | 
    72  | 
  | 
| 
 | 
    73  | 
(* *********************************************************************** *)
  | 
| 
 | 
    74  | 
(*                                                                         *)
  | 
| 
 | 
    75  | 
(* Abstracting data type                                                   *)
  | 
| 
 | 
    76  | 
(*                                                                         *)
  | 
| 
 | 
    77  | 
(* *********************************************************************** *)
  | 
| 
 | 
    78  | 
  | 
| 
 | 
    79  | 
(*Declaring the abstract list constructors*)
  | 
| 
 | 
    80  | 
  | 
| 
 | 
    81  | 
no_translations
  | 
| 
 | 
    82  | 
  "[x, xs]" == "x#[xs]"
  | 
| 
 | 
    83  | 
  "[x]" == "x#[]"
  | 
| 
 | 
    84  | 
no_notation
  | 
| 
 | 
    85  | 
  Nil  ("[]") and
 | 
| 
 | 
    86  | 
  Cons (infixr "#" 65)
  | 
| 
 | 
    87  | 
  | 
| 
 | 
    88  | 
definition
  | 
| 
 | 
    89  | 
  Nil       :: "'a list"                               ("[]") where
 | 
| 
 | 
    90  | 
   "Nil  = Abs_List(NIL)"
  | 
| 
 | 
    91  | 
  | 
| 
 | 
    92  | 
definition
  | 
| 
 | 
    93  | 
  "Cons"       :: "['a, 'a list] => 'a list"           (infixr "#" 65) where
  | 
| 
 | 
    94  | 
   "x#xs = Abs_List(CONS (Leaf x)(Rep_List xs))"
  | 
| 
 | 
    95  | 
  | 
| 
 | 
    96  | 
definition
  | 
| 
 | 
    97  | 
  (* list Recursion -- the trancl is Essential; see list.ML *)
  | 
| 
 | 
    98  | 
  list_rec  :: "['a list, 'b, ['a, 'a list, 'b]=>'b] => 'b" where
  | 
| 
 | 
    99  | 
   "list_rec l c d =
  | 
| 
 | 
   100  | 
      List_rec(Rep_List l) c (%x y r. d(inv Leaf x)(Abs_List y) r)"
  | 
| 
 | 
   101  | 
  | 
| 
 | 
   102  | 
definition
  | 
| 
 | 
   103  | 
  list_case :: "['b, ['a, 'a list]=>'b, 'a list] => 'b" where
  | 
| 
 | 
   104  | 
   "list_case a f xs = list_rec xs a (%x xs r. f x xs)"
  | 
| 
 | 
   105  | 
  | 
| 
 | 
   106  | 
(* list Enumeration *)
  | 
| 
 | 
   107  | 
translations
  | 
| 
 | 
   108  | 
  "[x, xs]" == "x#[xs]"
  | 
| 
 | 
   109  | 
  "[x]"     == "x#[]"
  | 
| 
 | 
   110  | 
  | 
| 
 | 
   111  | 
  "case xs of [] => a | y#ys => b" == "CONST list_case(a, %y ys. b, xs)"
  | 
| 
 | 
   112  | 
  | 
| 
 | 
   113  | 
  | 
| 
 | 
   114  | 
(* *********************************************************************** *)
  | 
| 
 | 
   115  | 
(*                                                                         *)
  | 
| 
 | 
   116  | 
(* Generalized Map Functionals                                             *)
  | 
| 
 | 
   117  | 
(*                                                                         *)
  | 
| 
 | 
   118  | 
(* *********************************************************************** *)
  | 
| 
 | 
   119  | 
  | 
| 
 | 
   120  | 
  
  | 
| 
 | 
   121  | 
(* Generalized Map Functionals *)
  | 
| 
 | 
   122  | 
  | 
| 
 | 
   123  | 
definition
  | 
| 
 | 
   124  | 
  Rep_map   :: "('b => 'a item) => ('b list => 'a item)" where
 | 
| 
 | 
   125  | 
   "Rep_map f xs = list_rec xs  NIL(%x l r. CONS(f x) r)"
  | 
| 
 | 
   126  | 
  | 
| 
 | 
   127  | 
definition
  | 
| 
 | 
   128  | 
  Abs_map   :: "('a item => 'b) => 'a item => 'b list" where
 | 
| 
 | 
   129  | 
   "Abs_map g M  = List_rec M Nil (%N L r. g(N)#r)"
  | 
| 
 | 
   130  | 
  | 
| 
 | 
   131  | 
  | 
| 
 | 
   132  | 
definition
  | 
| 
 | 
   133  | 
  map       :: "('a=>'b) => ('a list => 'b list)" where
 | 
| 
 | 
   134  | 
  "map f xs = list_rec xs [] (%x l r. f(x)#r)"
  | 
| 
 | 
   135  | 
  | 
| 
 | 
   136  | 
consts take      :: "['a list,nat] => 'a list"
  | 
| 
 | 
   137  | 
primrec
  | 
| 
 | 
   138  | 
  take_0:  "take xs 0 = []"
  | 
| 
 | 
   139  | 
  take_Suc: "take xs (Suc n) = list_case [] (%x l. x # take l n) xs"
  | 
| 
 | 
   140  | 
  | 
| 
 | 
   141  | 
lemma ListI: "x : list (range Leaf) ==> x : List"
  | 
| 
 | 
   142  | 
by (simp add: List_def)
  | 
| 
 | 
   143  | 
  | 
| 
 | 
   144  | 
lemma ListD: "x : List ==> x : list (range Leaf)"
  | 
| 
 | 
   145  | 
by (simp add: List_def)
  | 
| 
 | 
   146  | 
  | 
| 
 | 
   147  | 
lemma list_unfold: "list(A) = usum {Numb(0)} (uprod A (list(A)))"
 | 
| 
 | 
   148  | 
  by (fast intro!: list.intros [unfolded NIL_def CONS_def]
  | 
| 
 | 
   149  | 
           elim: list.cases [unfolded NIL_def CONS_def])
  | 
| 
 | 
   150  | 
  | 
| 
 | 
   151  | 
(*This justifies using list in other recursive type definitions*)
  | 
| 
 | 
   152  | 
lemma list_mono: "A<=B ==> list(A) <= list(B)"
  | 
| 
 | 
   153  | 
apply (rule subsetI)
  | 
| 
 | 
   154  | 
apply (erule list.induct)
  | 
| 
 | 
   155  | 
apply (auto intro!: list.intros)
  | 
| 
 | 
   156  | 
done
  | 
| 
 | 
   157  | 
  | 
| 
 | 
   158  | 
(*Type checking -- list creates well-founded sets*)
  | 
| 
 | 
   159  | 
lemma list_sexp: "list(sexp) <= sexp"
  | 
| 
 | 
   160  | 
apply (rule subsetI)
  | 
| 
 | 
   161  | 
apply (erule list.induct)
  | 
| 
 | 
   162  | 
apply (unfold NIL_def CONS_def)
  | 
| 
 | 
   163  | 
apply (auto intro: sexp.intros sexp_In0I sexp_In1I)
  | 
| 
 | 
   164  | 
done
  | 
| 
 | 
   165  | 
  | 
| 
 | 
   166  | 
(* A <= sexp ==> list(A) <= sexp *)
  | 
| 
 | 
   167  | 
lemmas list_subset_sexp = subset_trans [OF list_mono list_sexp] 
  | 
| 
 | 
   168  | 
  | 
| 
 | 
   169  | 
  | 
| 
 | 
   170  | 
(*Induction for the type 'a list *)
  | 
| 
 | 
   171  | 
lemma list_induct:
  | 
| 
 | 
   172  | 
    "[| P(Nil);    
  | 
| 
 | 
   173  | 
        !!x xs. P(xs) ==> P(x # xs) |]  ==> P(l)"
  | 
| 
 | 
   174  | 
apply (unfold Nil_def Cons_def) 
  | 
| 
 | 
   175  | 
apply (rule Rep_List_inverse [THEN subst])
  | 
| 
 | 
   176  | 
(*types force good instantiation*)
  | 
| 
 | 
   177  | 
apply (rule Rep_List [unfolded List_def, THEN list.induct], simp)
  | 
| 
 | 
   178  | 
apply (erule Abs_List_inverse [unfolded List_def, THEN subst], blast) 
  | 
| 
 | 
   179  | 
done
  | 
| 
 | 
   180  | 
  | 
| 
 | 
   181  | 
  | 
| 
 | 
   182  | 
(*** Isomorphisms ***)
  | 
| 
 | 
   183  | 
  | 
| 
 | 
   184  | 
lemma inj_on_Abs_list: "inj_on Abs_List (list(range Leaf))"
  | 
| 
 | 
   185  | 
apply (rule inj_on_inverseI)
  | 
| 
 | 
   186  | 
apply (erule Abs_List_inverse [unfolded List_def])
  | 
| 
 | 
   187  | 
done
  | 
| 
 | 
   188  | 
  | 
| 
 | 
   189  | 
(** Distinctness of constructors **)
  | 
| 
 | 
   190  | 
  | 
| 
 | 
   191  | 
lemma CONS_not_NIL [iff]: "CONS M N ~= NIL"
  | 
| 
 | 
   192  | 
by (simp add: NIL_def CONS_def)
  | 
| 
 | 
   193  | 
  | 
| 
 | 
   194  | 
lemmas NIL_not_CONS [iff] = CONS_not_NIL [THEN not_sym]
  | 
| 
 | 
   195  | 
lemmas CONS_neq_NIL = CONS_not_NIL [THEN notE, standard]
  | 
| 
 | 
   196  | 
lemmas NIL_neq_CONS = sym [THEN CONS_neq_NIL]
  | 
| 
 | 
   197  | 
  | 
| 
 | 
   198  | 
lemma Cons_not_Nil [iff]: "x # xs ~= Nil"
  | 
| 
 | 
   199  | 
apply (unfold Nil_def Cons_def)
  | 
| 
 | 
   200  | 
apply (rule CONS_not_NIL [THEN inj_on_Abs_list [THEN inj_on_contraD]])
  | 
| 
 | 
   201  | 
apply (simp_all add: list.intros rangeI Rep_List [unfolded List_def])
  | 
| 
 | 
   202  | 
done
  | 
| 
 | 
   203  | 
  | 
| 
 | 
   204  | 
lemmas Nil_not_Cons [iff] = Cons_not_Nil [THEN not_sym, standard]
  | 
| 
 | 
   205  | 
lemmas Cons_neq_Nil = Cons_not_Nil [THEN notE, standard]
  | 
| 
 | 
   206  | 
lemmas Nil_neq_Cons = sym [THEN Cons_neq_Nil]
  | 
| 
 | 
   207  | 
  | 
| 
 | 
   208  | 
(** Injectiveness of CONS and Cons **)
  | 
| 
 | 
   209  | 
  | 
| 
 | 
   210  | 
lemma CONS_CONS_eq [iff]: "(CONS K M)=(CONS L N) = (K=L & M=N)"
  | 
| 
 | 
   211  | 
by (simp add: CONS_def)
  | 
| 
 | 
   212  | 
  | 
| 
 | 
   213  | 
(*For reasoning about abstract list constructors*)
  | 
| 
 | 
   214  | 
declare Rep_List [THEN ListD, intro] ListI [intro]
  | 
| 
 | 
   215  | 
declare list.intros [intro,simp]
  | 
| 
 | 
   216  | 
declare Leaf_inject [dest!]
  | 
| 
 | 
   217  | 
  | 
| 
 | 
   218  | 
lemma Cons_Cons_eq [iff]: "(x#xs=y#ys) = (x=y & xs=ys)"
  | 
| 
 | 
   219  | 
apply (simp add: Cons_def)
  | 
| 
 | 
   220  | 
apply (subst Abs_List_inject)
  | 
| 
 | 
   221  | 
apply (auto simp add: Rep_List_inject)
  | 
| 
 | 
   222  | 
done
  | 
| 
 | 
   223  | 
  | 
| 
 | 
   224  | 
lemmas Cons_inject2 = Cons_Cons_eq [THEN iffD1, THEN conjE, standard]
  | 
| 
 | 
   225  | 
  | 
| 
 | 
   226  | 
lemma CONS_D: "CONS M N: list(A) ==> M: A & N: list(A)"
  | 
| 
 | 
   227  | 
  by (induct L == "CONS M N" set: list) auto
  | 
| 
 | 
   228  | 
  | 
| 
 | 
   229  | 
lemma sexp_CONS_D: "CONS M N: sexp ==> M: sexp & N: sexp"
  | 
| 
 | 
   230  | 
apply (simp add: CONS_def In1_def)
  | 
| 
 | 
   231  | 
apply (fast dest!: Scons_D)
  | 
| 
 | 
   232  | 
done
  | 
| 
 | 
   233  | 
  | 
| 
 | 
   234  | 
  | 
| 
 | 
   235  | 
(*Reasoning about constructors and their freeness*)
  | 
| 
 | 
   236  | 
  | 
| 
 | 
   237  | 
  | 
| 
 | 
   238  | 
lemma not_CONS_self: "N: list(A) ==> !M. N ~= CONS M N"
  | 
| 
 | 
   239  | 
apply (erule list.induct) apply simp_all done
  | 
| 
 | 
   240  | 
  | 
| 
 | 
   241  | 
lemma not_Cons_self2: "\<forall>x. l ~= x#l"
  | 
| 
 | 
   242  | 
by (induct l rule: list_induct) simp_all
  | 
| 
 | 
   243  | 
  | 
| 
 | 
   244  | 
  | 
| 
 | 
   245  | 
lemma neq_Nil_conv2: "(xs ~= []) = (\<exists>y ys. xs = y#ys)"
  | 
| 
 | 
   246  | 
by (induct xs rule: list_induct) auto
  | 
| 
 | 
   247  | 
  | 
| 
 | 
   248  | 
(** Conversion rules for List_case: case analysis operator **)
  | 
| 
 | 
   249  | 
  | 
| 
 | 
   250  | 
lemma List_case_NIL [simp]: "List_case c h NIL = c"
  | 
| 
 | 
   251  | 
by (simp add: List_case_def NIL_def)
  | 
| 
 | 
   252  | 
  | 
| 
 | 
   253  | 
lemma List_case_CONS [simp]: "List_case c h (CONS M N) = h M N"
  | 
| 
 | 
   254  | 
by (simp add: List_case_def CONS_def)
  | 
| 
 | 
   255  | 
  | 
| 
 | 
   256  | 
  | 
| 
 | 
   257  | 
(*** List_rec -- by wf recursion on pred_sexp ***)
  | 
| 
 | 
   258  | 
  | 
| 
 | 
   259  | 
(* The trancl(pred_sexp) is essential because pred_sexp_CONS_I1,2 would not
  | 
| 
 | 
   260  | 
   hold if pred_sexp^+ were changed to pred_sexp. *)
  | 
| 
 | 
   261  | 
  | 
| 
 | 
   262  | 
lemma List_rec_unfold_lemma:
  | 
| 
 | 
   263  | 
     "(%M. List_rec M c d) == 
  | 
| 
 | 
   264  | 
      wfrec (pred_sexp^+) (%g. List_case c (%x y. d x y (g y)))"
  | 
| 
 | 
   265  | 
by (simp add: List_rec_def)
  | 
| 
 | 
   266  | 
  | 
| 
 | 
   267  | 
lemmas List_rec_unfold = 
  | 
| 
 | 
   268  | 
    def_wfrec [OF List_rec_unfold_lemma wf_pred_sexp [THEN wf_trancl], 
  | 
| 
 | 
   269  | 
               standard]
  | 
| 
 | 
   270  | 
  | 
| 
 | 
   271  | 
  | 
| 
 | 
   272  | 
(** pred_sexp lemmas **)
  | 
| 
 | 
   273  | 
  | 
| 
 | 
   274  | 
lemma pred_sexp_CONS_I1: 
  | 
| 
 | 
   275  | 
    "[| M: sexp;  N: sexp |] ==> (M, CONS M N) : pred_sexp^+"
  | 
| 
 | 
   276  | 
by (simp add: CONS_def In1_def)
  | 
| 
 | 
   277  | 
  | 
| 
 | 
   278  | 
lemma pred_sexp_CONS_I2: 
  | 
| 
 | 
   279  | 
    "[| M: sexp;  N: sexp |] ==> (N, CONS M N) : pred_sexp^+"
  | 
| 
 | 
   280  | 
by (simp add: CONS_def In1_def)
  | 
| 
 | 
   281  | 
  | 
| 
 | 
   282  | 
lemma pred_sexp_CONS_D: 
  | 
| 
 | 
   283  | 
    "(CONS M1 M2, N) : pred_sexp^+ ==>  
  | 
| 
 | 
   284  | 
     (M1,N) : pred_sexp^+ & (M2,N) : pred_sexp^+"
  | 
| 
 | 
   285  | 
apply (frule pred_sexp_subset_Sigma [THEN trancl_subset_Sigma, THEN subsetD])
  | 
| 
 | 
   286  | 
apply (blast dest!: sexp_CONS_D intro: pred_sexp_CONS_I1 pred_sexp_CONS_I2 
  | 
| 
 | 
   287  | 
                    trans_trancl [THEN transD])
  | 
| 
 | 
   288  | 
done
  | 
| 
 | 
   289  | 
  | 
| 
 | 
   290  | 
  | 
| 
 | 
   291  | 
(** Conversion rules for List_rec **)
  | 
| 
 | 
   292  | 
  | 
| 
 | 
   293  | 
lemma List_rec_NIL [simp]: "List_rec NIL c h = c"
  | 
| 
 | 
   294  | 
apply (rule List_rec_unfold [THEN trans])
  | 
| 
 | 
   295  | 
apply (simp add: List_case_NIL)
  | 
| 
 | 
   296  | 
done
  | 
| 
 | 
   297  | 
  | 
| 
 | 
   298  | 
lemma List_rec_CONS [simp]:
  | 
| 
 | 
   299  | 
     "[| M: sexp;  N: sexp |] 
  | 
| 
 | 
   300  | 
      ==> List_rec (CONS M N) c h = h M N (List_rec N c h)"
  | 
| 
 | 
   301  | 
apply (rule List_rec_unfold [THEN trans])
  | 
| 
 | 
   302  | 
apply (simp add: pred_sexp_CONS_I2)
  | 
| 
 | 
   303  | 
done
  | 
| 
 | 
   304  | 
  | 
| 
 | 
   305  | 
  | 
| 
 | 
   306  | 
(*** list_rec -- by List_rec ***)
  | 
| 
 | 
   307  | 
  | 
| 
 | 
   308  | 
lemmas Rep_List_in_sexp =
  | 
| 
 | 
   309  | 
    subsetD [OF range_Leaf_subset_sexp [THEN list_subset_sexp]
  | 
| 
 | 
   310  | 
                Rep_List [THEN ListD]] 
  | 
| 
 | 
   311  | 
  | 
| 
 | 
   312  | 
  | 
| 
 | 
   313  | 
lemma list_rec_Nil [simp]: "list_rec Nil c h = c"
  | 
| 
 | 
   314  | 
by (simp add: list_rec_def ListI [THEN Abs_List_inverse] Nil_def)
  | 
| 
 | 
   315  | 
  | 
| 
 | 
   316  | 
  | 
| 
 | 
   317  | 
lemma list_rec_Cons [simp]: "list_rec (a#l) c h = h a l (list_rec l c h)"
  | 
| 
 | 
   318  | 
by (simp add: list_rec_def ListI [THEN Abs_List_inverse] Cons_def
  | 
| 
 | 
   319  | 
              Rep_List_inverse Rep_List [THEN ListD] inj_Leaf Rep_List_in_sexp)
  | 
| 
 | 
   320  | 
  | 
| 
 | 
   321  | 
  | 
| 
 | 
   322  | 
(*Type checking.  Useful?*)
  | 
| 
 | 
   323  | 
lemma List_rec_type:
  | 
| 
 | 
   324  | 
     "[| M: list(A);      
  | 
| 
 | 
   325  | 
         A<=sexp;         
  | 
| 
 | 
   326  | 
         c: C(NIL);       
  | 
| 
 | 
   327  | 
         !!x y r. [| x: A;  y: list(A);  r: C(y) |] ==> h x y r: C(CONS x y)  
  | 
| 
 | 
   328  | 
      |] ==> List_rec M c h : C(M :: 'a item)"
  | 
| 
 | 
   329  | 
apply (erule list.induct, simp) 
  | 
| 
 | 
   330  | 
apply (insert list_subset_sexp) 
  | 
| 
 | 
   331  | 
apply (subst List_rec_CONS, blast+)
  | 
| 
 | 
   332  | 
done
  | 
| 
 | 
   333  | 
  | 
| 
 | 
   334  | 
  | 
| 
 | 
   335  | 
  | 
| 
 | 
   336  | 
(** Generalized map functionals **)
  | 
| 
 | 
   337  | 
  | 
| 
 | 
   338  | 
lemma Rep_map_Nil [simp]: "Rep_map f Nil = NIL"
  | 
| 
 | 
   339  | 
by (simp add: Rep_map_def)
  | 
| 
 | 
   340  | 
  | 
| 
 | 
   341  | 
lemma Rep_map_Cons [simp]: 
  | 
| 
 | 
   342  | 
    "Rep_map f(x#xs) = CONS(f x)(Rep_map f xs)"
  | 
| 
 | 
   343  | 
by (simp add: Rep_map_def)
  | 
| 
 | 
   344  | 
  | 
| 
 | 
   345  | 
lemma Rep_map_type: "(!!x. f(x): A) ==> Rep_map f xs: list(A)"
  | 
| 
 | 
   346  | 
apply (simp add: Rep_map_def)
  | 
| 
 | 
   347  | 
apply (rule list_induct, auto)
  | 
| 
 | 
   348  | 
done
  | 
| 
 | 
   349  | 
  | 
| 
 | 
   350  | 
lemma Abs_map_NIL [simp]: "Abs_map g NIL = Nil"
  | 
| 
 | 
   351  | 
by (simp add: Abs_map_def)
  | 
| 
 | 
   352  | 
  | 
| 
 | 
   353  | 
lemma Abs_map_CONS [simp]: 
  | 
| 
 | 
   354  | 
    "[| M: sexp;  N: sexp |] ==> Abs_map g (CONS M N) = g(M) # Abs_map g N"
  | 
| 
 | 
   355  | 
by (simp add: Abs_map_def)
  | 
| 
 | 
   356  | 
  | 
| 
 | 
   357  | 
(*Eases the use of primitive recursion.*)
  | 
| 
 | 
   358  | 
lemma def_list_rec_NilCons:
  | 
| 
 | 
   359  | 
    "[| !!xs. f(xs) = list_rec xs c h  |] 
  | 
| 
 | 
   360  | 
     ==> f [] = c & f(x#xs) = h x xs (f xs)"
  | 
| 
 | 
   361  | 
by simp 
  | 
| 
 | 
   362  | 
  | 
| 
 | 
   363  | 
lemma Abs_map_inverse:
  | 
| 
 | 
   364  | 
     "[| M: list(A);  A<=sexp;  !!z. z: A ==> f(g(z)) = z |]  
  | 
| 
 | 
   365  | 
      ==> Rep_map f (Abs_map g M) = M"
  | 
| 
 | 
   366  | 
apply (erule list.induct, simp_all)
  | 
| 
 | 
   367  | 
apply (insert list_subset_sexp) 
  | 
| 
 | 
   368  | 
apply (subst Abs_map_CONS, blast)
  | 
| 
 | 
   369  | 
apply blast 
  | 
| 
 | 
   370  | 
apply simp 
  | 
| 
 | 
   371  | 
done
  | 
| 
 | 
   372  | 
  | 
| 
 | 
   373  | 
(*Rep_map_inverse is obtained via Abs_Rep_map and map_ident*)
  | 
| 
 | 
   374  | 
  | 
| 
 | 
   375  | 
(** list_case **)
  | 
| 
 | 
   376  | 
  | 
| 
 | 
   377  | 
(* setting up rewrite sets *)
  | 
| 
 | 
   378  | 
  | 
| 
 | 
   379  | 
text{*Better to have a single theorem with a conjunctive conclusion.*}
 | 
| 
 | 
   380  | 
declare def_list_rec_NilCons [OF list_case_def, simp]
  | 
| 
 | 
   381  | 
  | 
| 
 | 
   382  | 
(** list_case **)
  | 
| 
 | 
   383  | 
  | 
| 
 | 
   384  | 
lemma expand_list_case: 
  | 
| 
 | 
   385  | 
 "P(list_case a f xs) = ((xs=[] --> P a ) & (!y ys. xs=y#ys --> P(f y ys)))"
  | 
| 
 | 
   386  | 
by (induct xs rule: list_induct) simp_all
  | 
| 
 | 
   387  | 
  | 
| 
 | 
   388  | 
  | 
| 
 | 
   389  | 
(**** Function definitions ****)
  | 
| 
 | 
   390  | 
  | 
| 
 | 
   391  | 
declare def_list_rec_NilCons [OF map_def, simp]
  | 
| 
 | 
   392  | 
  | 
| 
 | 
   393  | 
(** The functional "map" and the generalized functionals **)
  | 
| 
 | 
   394  | 
  | 
| 
 | 
   395  | 
lemma Abs_Rep_map: 
  | 
| 
 | 
   396  | 
     "(!!x. f(x): sexp) ==>  
  | 
| 
 | 
   397  | 
        Abs_map g (Rep_map f xs) = map (%t. g(f(t))) xs"
  | 
| 
 | 
   398  | 
apply (induct xs rule: list_induct)
  | 
| 
 | 
   399  | 
apply (simp_all add: Rep_map_type list_sexp [THEN subsetD])
  | 
| 
 | 
   400  | 
done
  | 
| 
 | 
   401  | 
  | 
| 
 | 
   402  | 
  | 
| 
 | 
   403  | 
(** Additional mapping lemmas **)
  | 
| 
 | 
   404  | 
  | 
| 
 | 
   405  | 
lemma map_ident [simp]: "map(%x. x)(xs) = xs"
  | 
| 
 | 
   406  | 
by (induct xs rule: list_induct) simp_all
  | 
| 
 | 
   407  | 
  | 
| 
 | 
   408  | 
lemma map_compose: "map(f o g)(xs) = map f (map g xs)"
  | 
| 
 | 
   409  | 
apply (simp add: o_def)
  | 
| 
 | 
   410  | 
apply (induct xs rule: list_induct)
  | 
| 
 | 
   411  | 
apply simp_all
  | 
| 
 | 
   412  | 
done
  | 
| 
 | 
   413  | 
  | 
| 
 | 
   414  | 
  | 
| 
 | 
   415  | 
(** take **)
  | 
| 
 | 
   416  | 
  | 
| 
 | 
   417  | 
lemma take_Suc1 [simp]: "take [] (Suc x) = []"
  | 
| 
 | 
   418  | 
by simp
  | 
| 
 | 
   419  | 
  | 
| 
 | 
   420  | 
lemma take_Suc2 [simp]: "take(a#xs)(Suc x) = a#take xs x"
  | 
| 
 | 
   421  | 
by simp
  | 
| 
 | 
   422  | 
  | 
| 
 | 
   423  | 
lemma take_Nil [simp]: "take [] n = []"
  | 
| 
 | 
   424  | 
by (induct n) simp_all
  | 
| 
 | 
   425  | 
  | 
| 
 | 
   426  | 
lemma take_take_eq [simp]: "\<forall>n. take (take xs n) n = take xs n"
  | 
| 
 | 
   427  | 
apply (induct xs rule: list_induct)
  | 
| 
 | 
   428  | 
apply simp_all
  | 
| 
 | 
   429  | 
apply (rule allI)
  | 
| 
 | 
   430  | 
apply (induct_tac n)
  | 
| 
 | 
   431  | 
apply auto
  | 
| 
 | 
   432  | 
done
  | 
| 
 | 
   433  | 
  | 
| 
 | 
   434  | 
end
  |