| 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 | 
 | 
| 39246 |    136 | primrec take :: "['a list,nat] => 'a list" where
 | 
| 35247 |    137 |   take_0:  "take xs 0 = []"
 | 
| 39246 |    138 | | take_Suc: "take xs (Suc n) = list_case [] (%x l. x # take l n) xs"
 | 
| 35247 |    139 | 
 | 
|  |    140 | lemma ListI: "x : list (range Leaf) ==> x : List"
 | 
|  |    141 | by (simp add: List_def)
 | 
|  |    142 | 
 | 
|  |    143 | lemma ListD: "x : List ==> x : list (range Leaf)"
 | 
|  |    144 | by (simp add: List_def)
 | 
|  |    145 | 
 | 
|  |    146 | lemma list_unfold: "list(A) = usum {Numb(0)} (uprod A (list(A)))"
 | 
|  |    147 |   by (fast intro!: list.intros [unfolded NIL_def CONS_def]
 | 
|  |    148 |            elim: list.cases [unfolded NIL_def CONS_def])
 | 
|  |    149 | 
 | 
|  |    150 | (*This justifies using list in other recursive type definitions*)
 | 
|  |    151 | lemma list_mono: "A<=B ==> list(A) <= list(B)"
 | 
|  |    152 | apply (rule subsetI)
 | 
|  |    153 | apply (erule list.induct)
 | 
|  |    154 | apply (auto intro!: list.intros)
 | 
|  |    155 | done
 | 
|  |    156 | 
 | 
|  |    157 | (*Type checking -- list creates well-founded sets*)
 | 
|  |    158 | lemma list_sexp: "list(sexp) <= sexp"
 | 
|  |    159 | apply (rule subsetI)
 | 
|  |    160 | apply (erule list.induct)
 | 
|  |    161 | apply (unfold NIL_def CONS_def)
 | 
|  |    162 | apply (auto intro: sexp.intros sexp_In0I sexp_In1I)
 | 
|  |    163 | done
 | 
|  |    164 | 
 | 
|  |    165 | (* A <= sexp ==> list(A) <= sexp *)
 | 
|  |    166 | lemmas list_subset_sexp = subset_trans [OF list_mono list_sexp] 
 | 
|  |    167 | 
 | 
|  |    168 | 
 | 
|  |    169 | (*Induction for the type 'a list *)
 | 
|  |    170 | lemma list_induct:
 | 
|  |    171 |     "[| P(Nil);    
 | 
|  |    172 |         !!x xs. P(xs) ==> P(x # xs) |]  ==> P(l)"
 | 
|  |    173 | apply (unfold Nil_def Cons_def) 
 | 
|  |    174 | apply (rule Rep_List_inverse [THEN subst])
 | 
|  |    175 | (*types force good instantiation*)
 | 
|  |    176 | apply (rule Rep_List [unfolded List_def, THEN list.induct], simp)
 | 
|  |    177 | apply (erule Abs_List_inverse [unfolded List_def, THEN subst], blast) 
 | 
|  |    178 | done
 | 
|  |    179 | 
 | 
|  |    180 | 
 | 
|  |    181 | (*** Isomorphisms ***)
 | 
|  |    182 | 
 | 
|  |    183 | lemma inj_on_Abs_list: "inj_on Abs_List (list(range Leaf))"
 | 
|  |    184 | apply (rule inj_on_inverseI)
 | 
|  |    185 | apply (erule Abs_List_inverse [unfolded List_def])
 | 
|  |    186 | done
 | 
|  |    187 | 
 | 
|  |    188 | (** Distinctness of constructors **)
 | 
|  |    189 | 
 | 
|  |    190 | lemma CONS_not_NIL [iff]: "CONS M N ~= NIL"
 | 
|  |    191 | by (simp add: NIL_def CONS_def)
 | 
|  |    192 | 
 | 
|  |    193 | lemmas NIL_not_CONS [iff] = CONS_not_NIL [THEN not_sym]
 | 
|  |    194 | lemmas CONS_neq_NIL = CONS_not_NIL [THEN notE, standard]
 | 
|  |    195 | lemmas NIL_neq_CONS = sym [THEN CONS_neq_NIL]
 | 
|  |    196 | 
 | 
|  |    197 | lemma Cons_not_Nil [iff]: "x # xs ~= Nil"
 | 
|  |    198 | apply (unfold Nil_def Cons_def)
 | 
|  |    199 | apply (rule CONS_not_NIL [THEN inj_on_Abs_list [THEN inj_on_contraD]])
 | 
|  |    200 | apply (simp_all add: list.intros rangeI Rep_List [unfolded List_def])
 | 
|  |    201 | done
 | 
|  |    202 | 
 | 
|  |    203 | lemmas Nil_not_Cons [iff] = Cons_not_Nil [THEN not_sym, standard]
 | 
|  |    204 | lemmas Cons_neq_Nil = Cons_not_Nil [THEN notE, standard]
 | 
|  |    205 | lemmas Nil_neq_Cons = sym [THEN Cons_neq_Nil]
 | 
|  |    206 | 
 | 
|  |    207 | (** Injectiveness of CONS and Cons **)
 | 
|  |    208 | 
 | 
|  |    209 | lemma CONS_CONS_eq [iff]: "(CONS K M)=(CONS L N) = (K=L & M=N)"
 | 
|  |    210 | by (simp add: CONS_def)
 | 
|  |    211 | 
 | 
|  |    212 | (*For reasoning about abstract list constructors*)
 | 
|  |    213 | declare Rep_List [THEN ListD, intro] ListI [intro]
 | 
|  |    214 | declare list.intros [intro,simp]
 | 
|  |    215 | declare Leaf_inject [dest!]
 | 
|  |    216 | 
 | 
|  |    217 | lemma Cons_Cons_eq [iff]: "(x#xs=y#ys) = (x=y & xs=ys)"
 | 
|  |    218 | apply (simp add: Cons_def)
 | 
|  |    219 | apply (subst Abs_List_inject)
 | 
|  |    220 | apply (auto simp add: Rep_List_inject)
 | 
|  |    221 | done
 | 
|  |    222 | 
 | 
|  |    223 | lemmas Cons_inject2 = Cons_Cons_eq [THEN iffD1, THEN conjE, standard]
 | 
|  |    224 | 
 | 
|  |    225 | lemma CONS_D: "CONS M N: list(A) ==> M: A & N: list(A)"
 | 
|  |    226 |   by (induct L == "CONS M N" set: list) auto
 | 
|  |    227 | 
 | 
|  |    228 | lemma sexp_CONS_D: "CONS M N: sexp ==> M: sexp & N: sexp"
 | 
|  |    229 | apply (simp add: CONS_def In1_def)
 | 
|  |    230 | apply (fast dest!: Scons_D)
 | 
|  |    231 | done
 | 
|  |    232 | 
 | 
|  |    233 | 
 | 
|  |    234 | (*Reasoning about constructors and their freeness*)
 | 
|  |    235 | 
 | 
|  |    236 | 
 | 
|  |    237 | lemma not_CONS_self: "N: list(A) ==> !M. N ~= CONS M N"
 | 
|  |    238 | apply (erule list.induct) apply simp_all done
 | 
|  |    239 | 
 | 
|  |    240 | lemma not_Cons_self2: "\<forall>x. l ~= x#l"
 | 
|  |    241 | by (induct l rule: list_induct) simp_all
 | 
|  |    242 | 
 | 
|  |    243 | 
 | 
|  |    244 | lemma neq_Nil_conv2: "(xs ~= []) = (\<exists>y ys. xs = y#ys)"
 | 
|  |    245 | by (induct xs rule: list_induct) auto
 | 
|  |    246 | 
 | 
|  |    247 | (** Conversion rules for List_case: case analysis operator **)
 | 
|  |    248 | 
 | 
|  |    249 | lemma List_case_NIL [simp]: "List_case c h NIL = c"
 | 
|  |    250 | by (simp add: List_case_def NIL_def)
 | 
|  |    251 | 
 | 
|  |    252 | lemma List_case_CONS [simp]: "List_case c h (CONS M N) = h M N"
 | 
|  |    253 | by (simp add: List_case_def CONS_def)
 | 
|  |    254 | 
 | 
|  |    255 | 
 | 
|  |    256 | (*** List_rec -- by wf recursion on pred_sexp ***)
 | 
|  |    257 | 
 | 
|  |    258 | (* The trancl(pred_sexp) is essential because pred_sexp_CONS_I1,2 would not
 | 
|  |    259 |    hold if pred_sexp^+ were changed to pred_sexp. *)
 | 
|  |    260 | 
 | 
|  |    261 | lemma List_rec_unfold_lemma:
 | 
|  |    262 |      "(%M. List_rec M c d) == 
 | 
|  |    263 |       wfrec (pred_sexp^+) (%g. List_case c (%x y. d x y (g y)))"
 | 
|  |    264 | by (simp add: List_rec_def)
 | 
|  |    265 | 
 | 
|  |    266 | lemmas List_rec_unfold = 
 | 
|  |    267 |     def_wfrec [OF List_rec_unfold_lemma wf_pred_sexp [THEN wf_trancl], 
 | 
|  |    268 |                standard]
 | 
|  |    269 | 
 | 
|  |    270 | 
 | 
|  |    271 | (** pred_sexp lemmas **)
 | 
|  |    272 | 
 | 
|  |    273 | lemma pred_sexp_CONS_I1: 
 | 
|  |    274 |     "[| M: sexp;  N: sexp |] ==> (M, CONS M N) : pred_sexp^+"
 | 
|  |    275 | by (simp add: CONS_def In1_def)
 | 
|  |    276 | 
 | 
|  |    277 | lemma pred_sexp_CONS_I2: 
 | 
|  |    278 |     "[| M: sexp;  N: sexp |] ==> (N, CONS M N) : pred_sexp^+"
 | 
|  |    279 | by (simp add: CONS_def In1_def)
 | 
|  |    280 | 
 | 
|  |    281 | lemma pred_sexp_CONS_D: 
 | 
|  |    282 |     "(CONS M1 M2, N) : pred_sexp^+ ==>  
 | 
|  |    283 |      (M1,N) : pred_sexp^+ & (M2,N) : pred_sexp^+"
 | 
|  |    284 | apply (frule pred_sexp_subset_Sigma [THEN trancl_subset_Sigma, THEN subsetD])
 | 
|  |    285 | apply (blast dest!: sexp_CONS_D intro: pred_sexp_CONS_I1 pred_sexp_CONS_I2 
 | 
|  |    286 |                     trans_trancl [THEN transD])
 | 
|  |    287 | done
 | 
|  |    288 | 
 | 
|  |    289 | 
 | 
|  |    290 | (** Conversion rules for List_rec **)
 | 
|  |    291 | 
 | 
|  |    292 | lemma List_rec_NIL [simp]: "List_rec NIL c h = c"
 | 
|  |    293 | apply (rule List_rec_unfold [THEN trans])
 | 
|  |    294 | apply (simp add: List_case_NIL)
 | 
|  |    295 | done
 | 
|  |    296 | 
 | 
|  |    297 | lemma List_rec_CONS [simp]:
 | 
|  |    298 |      "[| M: sexp;  N: sexp |] 
 | 
|  |    299 |       ==> List_rec (CONS M N) c h = h M N (List_rec N c h)"
 | 
|  |    300 | apply (rule List_rec_unfold [THEN trans])
 | 
|  |    301 | apply (simp add: pred_sexp_CONS_I2)
 | 
|  |    302 | done
 | 
|  |    303 | 
 | 
|  |    304 | 
 | 
|  |    305 | (*** list_rec -- by List_rec ***)
 | 
|  |    306 | 
 | 
|  |    307 | lemmas Rep_List_in_sexp =
 | 
|  |    308 |     subsetD [OF range_Leaf_subset_sexp [THEN list_subset_sexp]
 | 
|  |    309 |                 Rep_List [THEN ListD]] 
 | 
|  |    310 | 
 | 
|  |    311 | 
 | 
|  |    312 | lemma list_rec_Nil [simp]: "list_rec Nil c h = c"
 | 
|  |    313 | by (simp add: list_rec_def ListI [THEN Abs_List_inverse] Nil_def)
 | 
|  |    314 | 
 | 
|  |    315 | 
 | 
|  |    316 | lemma list_rec_Cons [simp]: "list_rec (a#l) c h = h a l (list_rec l c h)"
 | 
|  |    317 | by (simp add: list_rec_def ListI [THEN Abs_List_inverse] Cons_def
 | 
|  |    318 |               Rep_List_inverse Rep_List [THEN ListD] inj_Leaf Rep_List_in_sexp)
 | 
|  |    319 | 
 | 
|  |    320 | 
 | 
|  |    321 | (*Type checking.  Useful?*)
 | 
|  |    322 | lemma List_rec_type:
 | 
|  |    323 |      "[| M: list(A);      
 | 
|  |    324 |          A<=sexp;         
 | 
|  |    325 |          c: C(NIL);       
 | 
|  |    326 |          !!x y r. [| x: A;  y: list(A);  r: C(y) |] ==> h x y r: C(CONS x y)  
 | 
|  |    327 |       |] ==> List_rec M c h : C(M :: 'a item)"
 | 
|  |    328 | apply (erule list.induct, simp) 
 | 
|  |    329 | apply (insert list_subset_sexp) 
 | 
|  |    330 | apply (subst List_rec_CONS, blast+)
 | 
|  |    331 | done
 | 
|  |    332 | 
 | 
|  |    333 | 
 | 
|  |    334 | 
 | 
|  |    335 | (** Generalized map functionals **)
 | 
|  |    336 | 
 | 
|  |    337 | lemma Rep_map_Nil [simp]: "Rep_map f Nil = NIL"
 | 
|  |    338 | by (simp add: Rep_map_def)
 | 
|  |    339 | 
 | 
|  |    340 | lemma Rep_map_Cons [simp]: 
 | 
|  |    341 |     "Rep_map f(x#xs) = CONS(f x)(Rep_map f xs)"
 | 
|  |    342 | by (simp add: Rep_map_def)
 | 
|  |    343 | 
 | 
|  |    344 | lemma Rep_map_type: "(!!x. f(x): A) ==> Rep_map f xs: list(A)"
 | 
|  |    345 | apply (simp add: Rep_map_def)
 | 
|  |    346 | apply (rule list_induct, auto)
 | 
|  |    347 | done
 | 
|  |    348 | 
 | 
|  |    349 | lemma Abs_map_NIL [simp]: "Abs_map g NIL = Nil"
 | 
|  |    350 | by (simp add: Abs_map_def)
 | 
|  |    351 | 
 | 
|  |    352 | lemma Abs_map_CONS [simp]: 
 | 
|  |    353 |     "[| M: sexp;  N: sexp |] ==> Abs_map g (CONS M N) = g(M) # Abs_map g N"
 | 
|  |    354 | by (simp add: Abs_map_def)
 | 
|  |    355 | 
 | 
|  |    356 | (*Eases the use of primitive recursion.*)
 | 
|  |    357 | lemma def_list_rec_NilCons:
 | 
|  |    358 |     "[| !!xs. f(xs) = list_rec xs c h  |] 
 | 
|  |    359 |      ==> f [] = c & f(x#xs) = h x xs (f xs)"
 | 
|  |    360 | by simp 
 | 
|  |    361 | 
 | 
|  |    362 | lemma Abs_map_inverse:
 | 
|  |    363 |      "[| M: list(A);  A<=sexp;  !!z. z: A ==> f(g(z)) = z |]  
 | 
|  |    364 |       ==> Rep_map f (Abs_map g M) = M"
 | 
|  |    365 | apply (erule list.induct, simp_all)
 | 
|  |    366 | apply (insert list_subset_sexp) 
 | 
|  |    367 | apply (subst Abs_map_CONS, blast)
 | 
|  |    368 | apply blast 
 | 
|  |    369 | apply simp 
 | 
|  |    370 | done
 | 
|  |    371 | 
 | 
|  |    372 | (*Rep_map_inverse is obtained via Abs_Rep_map and map_ident*)
 | 
|  |    373 | 
 | 
|  |    374 | (** list_case **)
 | 
|  |    375 | 
 | 
|  |    376 | (* setting up rewrite sets *)
 | 
|  |    377 | 
 | 
|  |    378 | text{*Better to have a single theorem with a conjunctive conclusion.*}
 | 
|  |    379 | declare def_list_rec_NilCons [OF list_case_def, simp]
 | 
|  |    380 | 
 | 
|  |    381 | (** list_case **)
 | 
|  |    382 | 
 | 
|  |    383 | lemma expand_list_case: 
 | 
|  |    384 |  "P(list_case a f xs) = ((xs=[] --> P a ) & (!y ys. xs=y#ys --> P(f y ys)))"
 | 
|  |    385 | by (induct xs rule: list_induct) simp_all
 | 
|  |    386 | 
 | 
|  |    387 | 
 | 
|  |    388 | (**** Function definitions ****)
 | 
|  |    389 | 
 | 
|  |    390 | declare def_list_rec_NilCons [OF map_def, simp]
 | 
|  |    391 | 
 | 
|  |    392 | (** The functional "map" and the generalized functionals **)
 | 
|  |    393 | 
 | 
|  |    394 | lemma Abs_Rep_map: 
 | 
|  |    395 |      "(!!x. f(x): sexp) ==>  
 | 
|  |    396 |         Abs_map g (Rep_map f xs) = map (%t. g(f(t))) xs"
 | 
|  |    397 | apply (induct xs rule: list_induct)
 | 
|  |    398 | apply (simp_all add: Rep_map_type list_sexp [THEN subsetD])
 | 
|  |    399 | done
 | 
|  |    400 | 
 | 
|  |    401 | 
 | 
|  |    402 | (** Additional mapping lemmas **)
 | 
|  |    403 | 
 | 
|  |    404 | lemma map_ident [simp]: "map(%x. x)(xs) = xs"
 | 
|  |    405 | by (induct xs rule: list_induct) simp_all
 | 
|  |    406 | 
 | 
|  |    407 | lemma map_compose: "map(f o g)(xs) = map f (map g xs)"
 | 
|  |    408 | apply (simp add: o_def)
 | 
|  |    409 | apply (induct xs rule: list_induct)
 | 
|  |    410 | apply simp_all
 | 
|  |    411 | done
 | 
|  |    412 | 
 | 
|  |    413 | 
 | 
|  |    414 | (** take **)
 | 
|  |    415 | 
 | 
|  |    416 | lemma take_Suc1 [simp]: "take [] (Suc x) = []"
 | 
|  |    417 | by simp
 | 
|  |    418 | 
 | 
|  |    419 | lemma take_Suc2 [simp]: "take(a#xs)(Suc x) = a#take xs x"
 | 
|  |    420 | by simp
 | 
|  |    421 | 
 | 
|  |    422 | lemma take_Nil [simp]: "take [] n = []"
 | 
|  |    423 | by (induct n) simp_all
 | 
|  |    424 | 
 | 
|  |    425 | lemma take_take_eq [simp]: "\<forall>n. take (take xs n) n = take xs n"
 | 
|  |    426 | apply (induct xs rule: list_induct)
 | 
|  |    427 | apply simp_all
 | 
|  |    428 | apply (rule allI)
 | 
|  |    429 | apply (induct_tac n)
 | 
|  |    430 | apply auto
 | 
|  |    431 | done
 | 
|  |    432 | 
 | 
|  |    433 | end
 |