src/HOL/Library/Eval.thy
author haftmann
Tue, 27 Mar 2007 09:19:37 +0200
changeset 22527 84690fcd3db9
parent 22525 d929a900584c
child 22665 cf152ff55d16
permissions -rw-r--r--
fixed document preparation
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
     1
(*  Title:      HOL/Library/Eval.thy
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
     2
    ID:         $Id$
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
     3
    Author:     Florian Haftmann, TU Muenchen
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
     4
*)
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
     5
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
     6
header {* A simple term evaluation mechanism *}
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
     7
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
     8
theory Eval
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
     9
imports Pure_term
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    10
begin
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    11
22527
84690fcd3db9 fixed document preparation
haftmann
parents: 22525
diff changeset
    12
subsection {* @{text typ_of} class *}
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    13
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    14
class typ_of = type +
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    15
  fixes typ_of :: "'a itself \<Rightarrow> typ"
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    16
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    17
ML {*
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    18
structure TypOf =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    19
struct
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    20
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    21
val class_typ_of = Sign.intern_class @{theory} "typ_of";
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    22
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    23
fun term_typ_of_type ty =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    24
  Const (@{const_name typ_of}, Term.itselfT ty --> @{typ typ})
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    25
    $ Logic.mk_type ty;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    26
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    27
fun mk_typ_of_def ty =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    28
  let
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    29
    val lhs = Const (@{const_name typ_of}, Term.itselfT ty --> @{typ typ})
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    30
      $ Free ("x", Term.itselfT ty)
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    31
    val rhs = Pure_term.mk_typ (fn v => term_typ_of_type (TFree v)) ty
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    32
  in Logic.mk_equals (lhs, rhs) end;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    33
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    34
end;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    35
*}
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    36
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    37
setup {*
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    38
let
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    39
  fun mk arities _ thy =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    40
    (maps (fn (tyco, asorts, _) => [(("", []), TypOf.mk_typ_of_def
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    41
      (Type (tyco,
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    42
        map TFree (Name.names Name.context "'a" asorts))))]) arities, thy);
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    43
  fun hook specs =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    44
    DatatypeCodegen.prove_codetypes_arities (ClassPackage.intro_classes_tac [])
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    45
      (map (fn (tyco, (is_dt, _)) => (tyco, is_dt)) specs)
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    46
      [TypOf.class_typ_of] mk ((K o K) I)
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    47
in DatatypeCodegen.add_codetypes_hook_bootstrap hook end
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    48
*}
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    49
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    50
22527
84690fcd3db9 fixed document preparation
haftmann
parents: 22525
diff changeset
    51
subsection {* @{text term_of} class *}
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    52
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    53
class term_of = typ_of +
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    54
  constrains typ_of :: "'a itself \<Rightarrow> typ"
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    55
  fixes term_of :: "'a \<Rightarrow> term"
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    56
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    57
ML {*
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    58
structure TermOf =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    59
struct
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    60
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    61
local
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    62
  fun term_term_of ty =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    63
    Const (@{const_name term_of}, ty --> @{typ term});
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    64
in
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    65
  val class_term_of = Sign.intern_class @{theory} "term_of";
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    66
  fun mk_terms_of_defs vs (tyco, cs) =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    67
    let
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    68
      val dty = Type (tyco, map TFree vs);
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    69
      fun mk_eq c =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    70
        let
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    71
          val lhs : term = term_term_of dty $ c;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    72
          val rhs : term = Pure_term.mk_term
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    73
            (fn (v, ty) => term_term_of ty $ Free (v, ty))
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    74
            (Pure_term.mk_typ (fn (v, sort) => TypOf.term_typ_of_type (TFree (v, sort)))) c
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    75
        in
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    76
          HOLogic.mk_eq (lhs, rhs)
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    77
        end;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    78
    in map mk_eq cs end;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    79
  fun mk_term_of t =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    80
    term_term_of (Term.fastype_of t) $ t;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    81
end;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    82
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    83
end;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    84
*}
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    85
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    86
setup {*
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    87
let
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    88
  fun thy_note ((name, atts), thms) =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    89
    PureThy.add_thmss [((name, thms), atts)] #-> (fn [thms] => pair (name, thms));
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    90
  fun thy_def ((name, atts), t) =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    91
    PureThy.add_defs_i false [((name, t), atts)] #-> (fn [thm] => pair (name, thm));
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    92
  fun mk arities css thy =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    93
    let
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    94
      val (_, asorts, _) :: _ = arities;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    95
      val vs = Name.names Name.context "'a" asorts;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    96
      val defs = map (TermOf.mk_terms_of_defs vs) css;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    97
      val defs' = (map (pair ("", []) o ObjectLogic.ensure_propT thy) o flat) defs;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    98
    in
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    99
      thy
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   100
      |> PrimrecPackage.gen_primrec thy_note thy_def "" defs'
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   101
      |> snd
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   102
    end;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   103
  fun hook specs =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   104
    if (fst o hd) specs = (fst o dest_Type) @{typ typ} then I
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   105
    else
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   106
      DatatypeCodegen.prove_codetypes_arities (ClassPackage.intro_classes_tac [])
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   107
      (map (fn (tyco, (is_dt, _)) => (tyco, is_dt)) specs)
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   108
      [TermOf.class_term_of] ((K o K o pair) []) mk
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   109
in DatatypeCodegen.add_codetypes_hook_bootstrap hook end
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   110
*}
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   111
22527
84690fcd3db9 fixed document preparation
haftmann
parents: 22525
diff changeset
   112
text {* Disable for @{typ char}acters and @{typ ml_string}s *}
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   113
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   114
code_const "typ_of \<Colon> char itself \<Rightarrow> typ" and "term_of \<Colon> char \<Rightarrow> term"
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   115
  (SML "!((_); raise Fail \"typ'_of'_char\")"
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   116
    and "!((_); raise Fail \"term'_of'_char\")")
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   117
  (OCaml "!((_); failwith \"typ'_of'_char\")"
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   118
    and "!((_); failwith \"term'_of'_char\")")
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   119
  (Haskell "error/ \"typ'_of'_char\""
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   120
    and "error/ \"term'_of'_char\"")
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   121
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   122
code_const "term_of \<Colon> ml_string \<Rightarrow> term"
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   123
  (SML "!((_); raise Fail \"term'_of'_ml'_string\")")
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   124
  (OCaml "!((_); failwith \"term'_of'_ml'_string\")")
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   125
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   126
subsection {* Evaluation infrastructure *}
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   127
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   128
ML {*
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   129
signature EVAL =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   130
sig
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   131
  val eval_term: theory -> term -> term
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   132
  val term: string -> unit
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   133
  val eval_ref: term option ref
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   134
end;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   135
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   136
structure Eval : EVAL =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   137
struct
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   138
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   139
val eval_ref = ref (NONE : term option);
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   140
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   141
fun eval_term thy t =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   142
  CodegenPackage.eval_term
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   143
    thy (("Eval.eval_ref", eval_ref), TermOf.mk_term_of t);
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   144
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   145
fun term t =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   146
  let
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   147
    val thy = the_context ();
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   148
    val t = eval_term thy (Sign.read_term thy t);
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   149
  in (writeln o Sign.string_of_term thy) t end;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   150
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   151
end;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   152
*}
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   153
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   154
end