src/HOL/List.thy
changeset 31048 ac146fc38b51
parent 31022 a438b4516dd3
child 31055 2cf6efca6c71
equal deleted inserted replaced
31047:c13b0406c039 31048:ac146fc38b51
     4 
     4 
     5 header {* The datatype of finite lists *}
     5 header {* The datatype of finite lists *}
     6 
     6 
     7 theory List
     7 theory List
     8 imports Plain Presburger Recdef ATP_Linkup
     8 imports Plain Presburger Recdef ATP_Linkup
     9 uses "Tools/string_syntax.ML"
       
    10 begin
     9 begin
    11 
    10 
    12 datatype 'a list =
    11 datatype 'a list =
    13     Nil    ("[]")
    12     Nil    ("[]")
    14   | Cons 'a  "'a list"    (infixr "#" 65)
    13   | Cons 'a  "'a list"    (infixr "#" 65)
  3431 lemma listrel_Cons:
  3430 lemma listrel_Cons:
  3432      "listrel r `` {x#xs} = set_Cons (r``{x}) (listrel r `` {xs})";
  3431      "listrel r `` {x#xs} = set_Cons (r``{x}) (listrel r `` {xs})";
  3433 by (auto simp add: set_Cons_def intro: listrel.intros) 
  3432 by (auto simp add: set_Cons_def intro: listrel.intros) 
  3434 
  3433 
  3435 
  3434 
  3436 subsection{*Miscellany*}
       
  3437 
       
  3438 subsubsection {* Characters and strings *}
       
  3439 
       
  3440 datatype nibble =
       
  3441     Nibble0 | Nibble1 | Nibble2 | Nibble3 | Nibble4 | Nibble5 | Nibble6 | Nibble7
       
  3442   | Nibble8 | Nibble9 | NibbleA | NibbleB | NibbleC | NibbleD | NibbleE | NibbleF
       
  3443 
       
  3444 lemma UNIV_nibble:
       
  3445   "UNIV = {Nibble0, Nibble1, Nibble2, Nibble3, Nibble4, Nibble5, Nibble6, Nibble7,
       
  3446     Nibble8, Nibble9, NibbleA, NibbleB, NibbleC, NibbleD, NibbleE, NibbleF}" (is "_ = ?A")
       
  3447 proof (rule UNIV_eq_I)
       
  3448   fix x show "x \<in> ?A" by (cases x) simp_all
       
  3449 qed
       
  3450 
       
  3451 instance nibble :: finite
       
  3452   by default (simp add: UNIV_nibble)
       
  3453 
       
  3454 datatype char = Char nibble nibble
       
  3455   -- "Note: canonical order of character encoding coincides with standard term ordering"
       
  3456 
       
  3457 lemma UNIV_char:
       
  3458   "UNIV = image (split Char) (UNIV \<times> UNIV)"
       
  3459 proof (rule UNIV_eq_I)
       
  3460   fix x show "x \<in> image (split Char) (UNIV \<times> UNIV)" by (cases x) auto
       
  3461 qed
       
  3462 
       
  3463 instance char :: finite
       
  3464   by default (simp add: UNIV_char)
       
  3465 
       
  3466 lemma size_char [code, simp]:
       
  3467   "size (c::char) = 0" by (cases c) simp
       
  3468 
       
  3469 lemma char_size [code, simp]:
       
  3470   "char_size (c::char) = 0" by (cases c) simp
       
  3471 
       
  3472 primrec nibble_pair_of_char :: "char \<Rightarrow> nibble \<times> nibble" where
       
  3473   "nibble_pair_of_char (Char n m) = (n, m)"
       
  3474 
       
  3475 declare nibble_pair_of_char.simps [code del]
       
  3476 
       
  3477 setup {*
       
  3478 let
       
  3479   val nibbles = map (Thm.cterm_of @{theory} o HOLogic.mk_nibble) (0 upto 15);
       
  3480   val thms = map_product
       
  3481    (fn n => fn m => Drule.instantiate' [] [SOME n, SOME m] @{thm nibble_pair_of_char.simps})
       
  3482       nibbles nibbles;
       
  3483 in
       
  3484   PureThy.note_thmss Thm.lemmaK [((Binding.name "nibble_pair_of_char_simps", []), [(thms, [])])]
       
  3485   #-> (fn [(_, thms)] => fold_rev Code.add_eqn thms)
       
  3486 end
       
  3487 *}
       
  3488 
       
  3489 lemma char_case_nibble_pair [code, code inline]:
       
  3490   "char_case f = split f o nibble_pair_of_char"
       
  3491   by (simp add: expand_fun_eq split: char.split)
       
  3492 
       
  3493 lemma char_rec_nibble_pair [code, code inline]:
       
  3494   "char_rec f = split f o nibble_pair_of_char"
       
  3495   unfolding char_case_nibble_pair [symmetric]
       
  3496   by (simp add: expand_fun_eq split: char.split)
       
  3497 
       
  3498 types string = "char list"
       
  3499 
       
  3500 syntax
       
  3501   "_Char" :: "xstr => char"    ("CHR _")
       
  3502   "_String" :: "xstr => string"    ("_")
       
  3503 
       
  3504 setup StringSyntax.setup
       
  3505 
       
  3506 
       
  3507 subsection {* Size function *}
  3435 subsection {* Size function *}
  3508 
  3436 
  3509 lemma [measure_function]: "is_measure f \<Longrightarrow> is_measure (list_size f)"
  3437 lemma [measure_function]: "is_measure f \<Longrightarrow> is_measure (list_size f)"
  3510 by (rule is_measure_trivial)
  3438 by (rule is_measure_trivial)
  3511 
  3439 
  3525 
  3453 
  3526 lemma list_size_pointwise[termination_simp]: 
  3454 lemma list_size_pointwise[termination_simp]: 
  3527   "(\<And>x. x \<in> set xs \<Longrightarrow> f x < g x) \<Longrightarrow> list_size f xs \<le> list_size g xs"
  3455   "(\<And>x. x \<in> set xs \<Longrightarrow> f x < g x) \<Longrightarrow> list_size f xs \<le> list_size g xs"
  3528 by (induct xs) force+
  3456 by (induct xs) force+
  3529 
  3457 
       
  3458 
  3530 subsection {* Code generator *}
  3459 subsection {* Code generator *}
  3531 
  3460 
  3532 subsubsection {* Setup *}
  3461 subsubsection {* Setup *}
       
  3462 
       
  3463 code_type list
       
  3464   (SML "_ list")
       
  3465   (OCaml "_ list")
       
  3466   (Haskell "![_]")
       
  3467 
       
  3468 code_const Nil
       
  3469   (SML "[]")
       
  3470   (OCaml "[]")
       
  3471   (Haskell "[]")
       
  3472 
       
  3473 code_const Cons
       
  3474   (SML infixr 7 "::")
       
  3475   (OCaml infixr 6 "::")
       
  3476   (Haskell infixr 5 ":")
       
  3477 
       
  3478 code_instance list :: eq
       
  3479   (Haskell -)
       
  3480 
       
  3481 code_const "eq_class.eq \<Colon> 'a\<Colon>eq list \<Rightarrow> 'a list \<Rightarrow> bool"
       
  3482   (Haskell infixl 4 "==")
       
  3483 
       
  3484 code_reserved SML
       
  3485   list
       
  3486 
       
  3487 code_reserved OCaml
       
  3488   list
  3533 
  3489 
  3534 types_code
  3490 types_code
  3535   "list" ("_ list")
  3491   "list" ("_ list")
  3536 attach (term_of) {*
  3492 attach (term_of) {*
  3537 fun term_of_list f T = HOLogic.mk_list T o map f;
  3493 fun term_of_list f T = HOLogic.mk_list T o map f;
  3544         val (xs, ts) = gen_list' aG aT (i-1) j
  3500         val (xs, ts) = gen_list' aG aT (i-1) j
  3545       in (x :: xs, fn () => HOLogic.cons_const aT $ t () $ ts ()) end),
  3501       in (x :: xs, fn () => HOLogic.cons_const aT $ t () $ ts ()) end),
  3546    (1, fn () => ([], fn () => HOLogic.nil_const aT))] ()
  3502    (1, fn () => ([], fn () => HOLogic.nil_const aT))] ()
  3547 and gen_list aG aT i = gen_list' aG aT i i;
  3503 and gen_list aG aT i = gen_list' aG aT i i;
  3548 *}
  3504 *}
  3549   "char" ("string")
  3505 
  3550 attach (term_of) {*
  3506 consts_code Nil ("[]")
  3551 val term_of_char = HOLogic.mk_char o ord;
  3507 consts_code Cons ("(_ ::/ _)")
  3552 *}
       
  3553 attach (test) {*
       
  3554 fun gen_char i =
       
  3555   let val j = random_range (ord "a") (Int.min (ord "a" + i, ord "z"))
       
  3556   in (chr j, fn () => HOLogic.mk_char j) end;
       
  3557 *}
       
  3558 
       
  3559 consts_code "Cons" ("(_ ::/ _)")
       
  3560 
       
  3561 code_type list
       
  3562   (SML "_ list")
       
  3563   (OCaml "_ list")
       
  3564   (Haskell "![_]")
       
  3565 
       
  3566 code_reserved SML
       
  3567   list
       
  3568 
       
  3569 code_reserved OCaml
       
  3570   list
       
  3571 
       
  3572 code_const Nil
       
  3573   (SML "[]")
       
  3574   (OCaml "[]")
       
  3575   (Haskell "[]")
       
  3576 
       
  3577 ML {*
       
  3578 local
       
  3579 
       
  3580 open Basic_Code_Thingol;
       
  3581 
       
  3582 fun implode_list naming t = case pairself
       
  3583   (Code_Thingol.lookup_const naming) (@{const_name Nil}, @{const_name Cons})
       
  3584    of (SOME nil', SOME cons') => let
       
  3585           fun dest_cons (IConst (c, _) `$ t1 `$ t2) =
       
  3586                 if c = cons'
       
  3587                 then SOME (t1, t2)
       
  3588                 else NONE
       
  3589             | dest_cons _ = NONE;
       
  3590           val (ts, t') = Code_Thingol.unfoldr dest_cons t;
       
  3591         in case t'
       
  3592          of IConst (c, _) => if c = nil' then SOME ts else NONE
       
  3593           | _ => NONE
       
  3594         end
       
  3595     | _ => NONE
       
  3596 
       
  3597 fun decode_char naming (IConst (c1, _), IConst (c2, _)) = (case map_filter
       
  3598   (Code_Thingol.lookup_const naming)[@{const_name Nibble0}, @{const_name Nibble1},
       
  3599    @{const_name Nibble2}, @{const_name Nibble3},
       
  3600    @{const_name Nibble4}, @{const_name Nibble5},
       
  3601    @{const_name Nibble6}, @{const_name Nibble7},
       
  3602    @{const_name Nibble8}, @{const_name Nibble9},
       
  3603    @{const_name NibbleA}, @{const_name NibbleB},
       
  3604    @{const_name NibbleC}, @{const_name NibbleD},
       
  3605    @{const_name NibbleE}, @{const_name NibbleF}]
       
  3606    of nibbles' as [_, _, _, _, _, _, _, _, _, _, _, _, _, _, _, _] => let
       
  3607           fun idx c = find_index (curry (op =) c) nibbles';
       
  3608           fun decode ~1 _ = NONE
       
  3609             | decode _ ~1 = NONE
       
  3610             | decode n m = SOME (chr (n * 16 + m));
       
  3611         in decode (idx c1) (idx c2) end
       
  3612     | _ => NONE)
       
  3613  | decode_char _ _ = NONE
       
  3614    
       
  3615 fun implode_string naming mk_char mk_string ts = case
       
  3616   Code_Thingol.lookup_const naming @{const_name Char}
       
  3617    of SOME char' => let
       
  3618         fun implode_char (IConst (c, _) `$ t1 `$ t2) =
       
  3619               if c = char' then decode_char naming (t1, t2) else NONE
       
  3620           | implode_char _ = NONE;
       
  3621         val ts' = map implode_char ts;
       
  3622       in if forall is_some ts'
       
  3623         then (SOME o Code_Printer.str o mk_string o implode o map_filter I) ts'
       
  3624         else NONE
       
  3625       end
       
  3626     | _ => NONE;
       
  3627 
       
  3628 fun default_list (target_fxy, target_cons) pr fxy t1 t2 =
       
  3629   Code_Printer.brackify_infix (target_fxy, Code_Printer.R) fxy [
       
  3630     pr (Code_Printer.INFX (target_fxy, Code_Printer.X)) t1,
       
  3631     Code_Printer.str target_cons,
       
  3632     pr (Code_Printer.INFX (target_fxy, Code_Printer.R)) t2
       
  3633   ];
       
  3634 
       
  3635 fun pretty_list literals =
       
  3636   let
       
  3637     val mk_list = Code_Printer.literal_list literals;
       
  3638     fun pretty pr naming thm vars fxy [(t1, _), (t2, _)] =
       
  3639       case Option.map (cons t1) (implode_list naming t2)
       
  3640        of SOME ts => mk_list (map (pr vars Code_Printer.NOBR) ts)
       
  3641         | NONE => default_list (Code_Printer.infix_cons literals) (pr vars) fxy t1 t2;
       
  3642   in (2, pretty) end;
       
  3643 
       
  3644 fun pretty_list_string literals =
       
  3645   let
       
  3646     val mk_list = Code_Printer.literal_list literals;
       
  3647     val mk_char = Code_Printer.literal_char literals;
       
  3648     val mk_string = Code_Printer.literal_string literals;
       
  3649     fun pretty pr naming thm vars fxy [(t1, _), (t2, _)] =
       
  3650       case Option.map (cons t1) (implode_list naming t2)
       
  3651        of SOME ts => (case implode_string naming mk_char mk_string ts
       
  3652            of SOME p => p
       
  3653             | NONE => mk_list (map (pr vars Code_Printer.NOBR) ts))
       
  3654         | NONE => default_list (Code_Printer.infix_cons literals) (pr vars) fxy t1 t2;
       
  3655   in (2, pretty) end;
       
  3656 
       
  3657 fun pretty_char literals =
       
  3658   let
       
  3659     val mk_char = Code_Printer.literal_char literals;
       
  3660     fun pretty _ naming thm _ _ [(t1, _), (t2, _)] =
       
  3661       case decode_char naming (t1, t2)
       
  3662        of SOME c => (Code_Printer.str o mk_char) c
       
  3663         | NONE => Code_Printer.nerror thm "Illegal character expression";
       
  3664   in (2, pretty) end;
       
  3665 
       
  3666 fun pretty_message literals =
       
  3667   let
       
  3668     val mk_char = Code_Printer.literal_char literals;
       
  3669     val mk_string = Code_Printer.literal_string literals;
       
  3670     fun pretty _ naming thm _ _ [(t, _)] =
       
  3671       case implode_list naming t
       
  3672        of SOME ts => (case implode_string naming mk_char mk_string ts
       
  3673            of SOME p => p
       
  3674             | NONE => Code_Printer.nerror thm "Illegal message expression")
       
  3675         | NONE => Code_Printer.nerror thm "Illegal message expression";
       
  3676   in (1, pretty) end;
       
  3677 
       
  3678 in
       
  3679 
       
  3680 fun add_literal_list target thy =
       
  3681   let
       
  3682     val pr = pretty_list (Code_Target.the_literals thy target);
       
  3683   in
       
  3684     thy
       
  3685     |> Code_Target.add_syntax_const target @{const_name Cons} (SOME pr)
       
  3686   end;
       
  3687 
       
  3688 fun add_literal_list_string target thy =
       
  3689   let
       
  3690     val pr = pretty_list_string (Code_Target.the_literals thy target);
       
  3691   in
       
  3692     thy
       
  3693     |> Code_Target.add_syntax_const target @{const_name Cons} (SOME pr)
       
  3694   end;
       
  3695 
       
  3696 fun add_literal_char target thy =
       
  3697   let
       
  3698     val pr = pretty_char (Code_Target.the_literals thy target);
       
  3699   in
       
  3700     thy
       
  3701     |> Code_Target.add_syntax_const target @{const_name Char} (SOME pr)
       
  3702   end;
       
  3703 
       
  3704 fun add_literal_message str target thy =
       
  3705   let
       
  3706     val pr = pretty_message (Code_Target.the_literals thy target);
       
  3707   in
       
  3708     thy
       
  3709     |> Code_Target.add_syntax_const target str (SOME pr)
       
  3710   end;
       
  3711 
       
  3712 end;
       
  3713 *}
       
  3714 
       
  3715 setup {*
       
  3716   fold (fn target => add_literal_list target) ["SML", "OCaml", "Haskell"]
       
  3717 *}
       
  3718 
       
  3719 code_instance list :: eq
       
  3720   (Haskell -)
       
  3721 
       
  3722 code_const "eq_class.eq \<Colon> 'a\<Colon>eq list \<Rightarrow> 'a list \<Rightarrow> bool"
       
  3723   (Haskell infixl 4 "==")
       
  3724 
  3508 
  3725 setup {*
  3509 setup {*
  3726 let
  3510 let
  3727 
  3511 
  3728 fun list_codegen thy defs dep thyname b t gr =
  3512 fun list_codegen thy defs dep thyname b t gr =
  3732       (fastype_of t) gr;
  3516       (fastype_of t) gr;
  3733     val (ps, gr'') = fold_map
  3517     val (ps, gr'') = fold_map
  3734       (Codegen.invoke_codegen thy defs dep thyname false) ts gr'
  3518       (Codegen.invoke_codegen thy defs dep thyname false) ts gr'
  3735   in SOME (Pretty.list "[" "]" ps, gr'') end handle TERM _ => NONE;
  3519   in SOME (Pretty.list "[" "]" ps, gr'') end handle TERM _ => NONE;
  3736 
  3520 
  3737 fun char_codegen thy defs dep thyname b t gr =
  3521 in Codegen.add_codegen "list_codegen" list_codegen end
  3738   let
       
  3739     val i = HOLogic.dest_char t;
       
  3740     val (_, gr') = Codegen.invoke_tycodegen thy defs dep thyname false
       
  3741       (fastype_of t) gr;
       
  3742   in SOME (Codegen.str (ML_Syntax.print_string (chr i)), gr')
       
  3743   end handle TERM _ => NONE;
       
  3744 
       
  3745 in
       
  3746   Codegen.add_codegen "list_codegen" list_codegen
       
  3747   #> Codegen.add_codegen "char_codegen" char_codegen
       
  3748 end;
       
  3749 *}
  3522 *}
  3750 
  3523 
  3751 
  3524 
  3752 subsubsection {* Generation of efficient code *}
  3525 subsubsection {* Generation of efficient code *}
  3753 
  3526