src/HOL/Library/Eval.thy
author haftmann
Thu, 06 Dec 2007 15:10:09 +0100
changeset 25557 ea6b11021e79
parent 25536 01753a944433
child 25559 f14305fb698c
permissions -rw-r--r--
added new primrec package
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
24280
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
     9
imports Main Pure_term
22525
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
23062
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
    14
class typ_of =
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
    15
  fixes typ_of :: "'a\<Colon>{} itself \<Rightarrow> typ"
22525
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
25536
01753a944433 improved
haftmann
parents: 25502
diff changeset
    21
fun mk ty =
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    22
  Const (@{const_name typ_of}, Term.itselfT ty --> @{typ typ})
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    23
    $ Logic.mk_type ty;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    24
25536
01753a944433 improved
haftmann
parents: 25502
diff changeset
    25
end
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    26
*}
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    27
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    28
setup {*
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    29
let
25536
01753a944433 improved
haftmann
parents: 25502
diff changeset
    30
  fun define_typ_of ty lthy =
01753a944433 improved
haftmann
parents: 25502
diff changeset
    31
    let
01753a944433 improved
haftmann
parents: 25502
diff changeset
    32
      val lhs = Const (@{const_name typ_of}, Term.itselfT ty --> @{typ typ})
01753a944433 improved
haftmann
parents: 25502
diff changeset
    33
        $ Free ("T", Term.itselfT ty);
01753a944433 improved
haftmann
parents: 25502
diff changeset
    34
      val rhs = Pure_term.mk_typ (fn v => TypOf.mk (TFree v)) ty;
01753a944433 improved
haftmann
parents: 25502
diff changeset
    35
      val eq = Class.prep_spec lthy (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)))
01753a944433 improved
haftmann
parents: 25502
diff changeset
    36
    in lthy |> Specification.definition (NONE, (("", []), eq)) end;
01753a944433 improved
haftmann
parents: 25502
diff changeset
    37
  fun interpretator tyco thy =
01753a944433 improved
haftmann
parents: 25502
diff changeset
    38
    let
01753a944433 improved
haftmann
parents: 25502
diff changeset
    39
      val sorts = replicate (Sign.arity_number thy tyco) @{sort typ_of};
01753a944433 improved
haftmann
parents: 25502
diff changeset
    40
      val ty = Type (tyco, map TFree (Name.names Name.context "'a" sorts));
01753a944433 improved
haftmann
parents: 25502
diff changeset
    41
    in
01753a944433 improved
haftmann
parents: 25502
diff changeset
    42
      thy
01753a944433 improved
haftmann
parents: 25502
diff changeset
    43
      |> Instance.instantiate ([tyco], sorts, @{sort typ_of})
01753a944433 improved
haftmann
parents: 25502
diff changeset
    44
           (define_typ_of ty) ((K o K) (Class.intro_classes_tac []))
01753a944433 improved
haftmann
parents: 25502
diff changeset
    45
    end;
01753a944433 improved
haftmann
parents: 25502
diff changeset
    46
in TypedefPackage.interpretation interpretator end
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    47
*}
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    48
25536
01753a944433 improved
haftmann
parents: 25502
diff changeset
    49
instantiation "prop" :: typ_of
01753a944433 improved
haftmann
parents: 25502
diff changeset
    50
begin
01753a944433 improved
haftmann
parents: 25502
diff changeset
    51
01753a944433 improved
haftmann
parents: 25502
diff changeset
    52
definition
01753a944433 improved
haftmann
parents: 25502
diff changeset
    53
  "typ_of (T\<Colon>prop itself) = STR ''prop'' {\<struct>} []"
01753a944433 improved
haftmann
parents: 25502
diff changeset
    54
01753a944433 improved
haftmann
parents: 25502
diff changeset
    55
instance ..
01753a944433 improved
haftmann
parents: 25502
diff changeset
    56
01753a944433 improved
haftmann
parents: 25502
diff changeset
    57
end
01753a944433 improved
haftmann
parents: 25502
diff changeset
    58
01753a944433 improved
haftmann
parents: 25502
diff changeset
    59
instantiation itself :: (typ_of) typ_of
01753a944433 improved
haftmann
parents: 25502
diff changeset
    60
begin
01753a944433 improved
haftmann
parents: 25502
diff changeset
    61
01753a944433 improved
haftmann
parents: 25502
diff changeset
    62
definition
01753a944433 improved
haftmann
parents: 25502
diff changeset
    63
  "typ_of (T\<Colon>'a itself itself) = STR ''itself'' {\<struct>} [typ_of TYPE('a\<Colon>typ_of)]"
01753a944433 improved
haftmann
parents: 25502
diff changeset
    64
01753a944433 improved
haftmann
parents: 25502
diff changeset
    65
instance ..
01753a944433 improved
haftmann
parents: 25502
diff changeset
    66
01753a944433 improved
haftmann
parents: 25502
diff changeset
    67
end
01753a944433 improved
haftmann
parents: 25502
diff changeset
    68
01753a944433 improved
haftmann
parents: 25502
diff changeset
    69
instantiation set :: (typ_of) typ_of
01753a944433 improved
haftmann
parents: 25502
diff changeset
    70
begin
01753a944433 improved
haftmann
parents: 25502
diff changeset
    71
 
01753a944433 improved
haftmann
parents: 25502
diff changeset
    72
definition
01753a944433 improved
haftmann
parents: 25502
diff changeset
    73
  "typ_of (T\<Colon>'a set itself) = STR ''set'' {\<struct>} [typ_of TYPE('a\<Colon>typ_of)]"
01753a944433 improved
haftmann
parents: 25502
diff changeset
    74
01753a944433 improved
haftmann
parents: 25502
diff changeset
    75
instance ..
01753a944433 improved
haftmann
parents: 25502
diff changeset
    76
01753a944433 improved
haftmann
parents: 25502
diff changeset
    77
end
01753a944433 improved
haftmann
parents: 25502
diff changeset
    78
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    79
22527
84690fcd3db9 fixed document preparation
haftmann
parents: 22525
diff changeset
    80
subsection {* @{text term_of} class *}
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    81
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    82
class term_of = typ_of +
23062
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
    83
  constrains typ_of :: "'a\<Colon>{} itself \<Rightarrow> typ"
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    84
  fixes term_of :: "'a \<Rightarrow> term"
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    85
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    86
ML {*
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    87
structure TermOf =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    88
struct
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    89
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    90
local
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    91
  fun term_term_of ty =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    92
    Const (@{const_name term_of}, ty --> @{typ term});
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    93
in
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    94
  val class_term_of = Sign.intern_class @{theory} "term_of";
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    95
  fun mk_terms_of_defs vs (tyco, cs) =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    96
    let
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    97
      val dty = Type (tyco, map TFree vs);
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    98
      fun mk_eq c =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    99
        let
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   100
          val lhs : term = term_term_of dty $ c;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   101
          val rhs : term = Pure_term.mk_term
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   102
            (fn (v, ty) => term_term_of ty $ Free (v, ty))
25536
01753a944433 improved
haftmann
parents: 25502
diff changeset
   103
            (Pure_term.mk_typ (fn (v, sort) => TypOf.mk (TFree (v, sort)))) c
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   104
        in
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   105
          HOLogic.mk_eq (lhs, rhs)
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   106
        end;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   107
    in map mk_eq cs end;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   108
  fun mk_term_of t =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   109
    term_term_of (Term.fastype_of t) $ t;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   110
end;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   111
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   112
end;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   113
*}
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   114
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   115
setup {*
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   116
let
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   117
  fun thy_note ((name, atts), thms) =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   118
    PureThy.add_thmss [((name, thms), atts)] #-> (fn [thms] => pair (name, thms));
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   119
  fun thy_def ((name, atts), t) =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   120
    PureThy.add_defs_i false [((name, t), atts)] #-> (fn [thm] => pair (name, thm));
25536
01753a944433 improved
haftmann
parents: 25502
diff changeset
   121
  fun prep_dtyp thy vs tyco =
01753a944433 improved
haftmann
parents: 25502
diff changeset
   122
    let
01753a944433 improved
haftmann
parents: 25502
diff changeset
   123
      val (_, cs) = DatatypePackage.the_datatype_spec thy tyco;
01753a944433 improved
haftmann
parents: 25502
diff changeset
   124
      val prep_typ = map_atyps (fn TFree (v, sort) =>
01753a944433 improved
haftmann
parents: 25502
diff changeset
   125
        TFree (v, (the o AList.lookup (op =) vs) v));
01753a944433 improved
haftmann
parents: 25502
diff changeset
   126
      fun prep_c (c, tys) = list_comb (Const (c, tys ---> Type (tyco, map TFree vs)),
01753a944433 improved
haftmann
parents: 25502
diff changeset
   127
        map Free (Name.names Name.context "a" tys));
01753a944433 improved
haftmann
parents: 25502
diff changeset
   128
    in (tyco, map (prep_c o (apsnd o map) prep_typ) cs) end;
01753a944433 improved
haftmann
parents: 25502
diff changeset
   129
  fun prep thy tycos =
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   130
    let
25536
01753a944433 improved
haftmann
parents: 25502
diff changeset
   131
      val inter_sort = curry (Sorts.inter_sort (Sign.classes_of thy)) @{sort term_of};
01753a944433 improved
haftmann
parents: 25502
diff changeset
   132
      val tyco = hd tycos;
01753a944433 improved
haftmann
parents: 25502
diff changeset
   133
      val (vs_proto, _) = DatatypePackage.the_datatype_spec thy tyco;
01753a944433 improved
haftmann
parents: 25502
diff changeset
   134
      val all_typs = maps (maps snd o snd o DatatypePackage.the_datatype_spec thy) tycos;
01753a944433 improved
haftmann
parents: 25502
diff changeset
   135
      fun add_tycos (Type (tyco, tys)) = insert (op =) tyco #>
01753a944433 improved
haftmann
parents: 25502
diff changeset
   136
            fold add_tycos tys
01753a944433 improved
haftmann
parents: 25502
diff changeset
   137
        | add_tycos _ = I;
01753a944433 improved
haftmann
parents: 25502
diff changeset
   138
      val dep_tycos = [] |> fold add_tycos all_typs |> subtract (op =) tycos;
01753a944433 improved
haftmann
parents: 25502
diff changeset
   139
      val sorts = map (inter_sort o snd) vs_proto;
01753a944433 improved
haftmann
parents: 25502
diff changeset
   140
      val vs = map fst vs_proto ~~ sorts;
01753a944433 improved
haftmann
parents: 25502
diff changeset
   141
      val css = map (prep_dtyp thy vs) tycos;
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   142
      val defs = map (TermOf.mk_terms_of_defs vs) css;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   143
      val defs' = (map (pair ("", []) o ObjectLogic.ensure_propT thy) o flat) defs;
25536
01753a944433 improved
haftmann
parents: 25502
diff changeset
   144
    in if forall (fn tyco => can (Sign.arity_sorts thy tyco) @{sort term_of}) dep_tycos
01753a944433 improved
haftmann
parents: 25502
diff changeset
   145
        andalso not (tycos = [@{type_name typ}])
01753a944433 improved
haftmann
parents: 25502
diff changeset
   146
      then SOME (sorts, defs')
01753a944433 improved
haftmann
parents: 25502
diff changeset
   147
      else NONE
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   148
    end;
25536
01753a944433 improved
haftmann
parents: 25502
diff changeset
   149
  fun interpretator tycos thy = case prep thy tycos
01753a944433 improved
haftmann
parents: 25502
diff changeset
   150
   of SOME (sorts, defs) =>
01753a944433 improved
haftmann
parents: 25502
diff changeset
   151
      thy
01753a944433 improved
haftmann
parents: 25502
diff changeset
   152
      |> Instance.instantiate (tycos, sorts, @{sort term_of})
01753a944433 improved
haftmann
parents: 25502
diff changeset
   153
           (pair ()) ((K o K) (Class.intro_classes_tac []))
25557
ea6b11021e79 added new primrec package
haftmann
parents: 25536
diff changeset
   154
      |> OldPrimrecPackage.gen_primrec thy_note thy_def "" defs
25536
01753a944433 improved
haftmann
parents: 25502
diff changeset
   155
      |> snd
01753a944433 improved
haftmann
parents: 25502
diff changeset
   156
    | NONE => thy;
01753a944433 improved
haftmann
parents: 25502
diff changeset
   157
  in DatatypePackage.interpretation interpretator end
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   158
*}
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   159
23062
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   160
abbreviation
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   161
  intT :: "typ"
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   162
where
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   163
  "intT \<equiv> STR ''IntDef.int'' {\<struct>} []"
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   164
23133
5a6935d598c3 fixed typo
haftmann
parents: 23062
diff changeset
   165
abbreviation
5a6935d598c3 fixed typo
haftmann
parents: 23062
diff changeset
   166
  bitT :: "typ"
5a6935d598c3 fixed typo
haftmann
parents: 23062
diff changeset
   167
where
5a6935d598c3 fixed typo
haftmann
parents: 23062
diff changeset
   168
  "bitT \<equiv> STR ''Numeral.bit'' {\<struct>} []"
5a6935d598c3 fixed typo
haftmann
parents: 23062
diff changeset
   169
23062
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   170
function
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   171
  mk_int :: "int \<Rightarrow> term"
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   172
where
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   173
  "mk_int k = (if k = 0 then STR ''Numeral.Pls'' \<Colon>\<subseteq> intT
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   174
    else if k = -1 then STR ''Numeral.Min'' \<Colon>\<subseteq> intT
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   175
    else let (l, m) = divAlg (k, 2)
23133
5a6935d598c3 fixed typo
haftmann
parents: 23062
diff changeset
   176
  in STR ''Numeral.Bit'' \<Colon>\<subseteq> intT \<rightarrow> bitT \<rightarrow> intT \<bullet> mk_int l \<bullet>
5a6935d598c3 fixed typo
haftmann
parents: 23062
diff changeset
   177
    (if m = 0 then STR ''Numeral.bit.B0'' \<Colon>\<subseteq> bitT else STR ''Numeral.bit.B1'' \<Colon>\<subseteq> bitT))"
23062
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   178
by pat_completeness auto
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   179
termination by (relation "measure (nat o abs)") (auto simp add: divAlg_mod_div)
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   180
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   181
instance int :: term_of
23133
5a6935d598c3 fixed typo
haftmann
parents: 23062
diff changeset
   182
  "term_of k \<equiv> STR ''Numeral.number_class.number_of'' \<Colon>\<subseteq> intT \<rightarrow> intT \<bullet> mk_int k" ..
23062
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   183
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   184
24994
c385c4eabb3b consolidated naming conventions for code generator theories
haftmann
parents: 24920
diff changeset
   185
text {* Adaption for @{typ message_string}s *}
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   186
24994
c385c4eabb3b consolidated naming conventions for code generator theories
haftmann
parents: 24920
diff changeset
   187
lemmas [code func, code func del] = term_of_message_string_def
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   188
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   189
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   190
subsection {* Evaluation infrastructure *}
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   191
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   192
ML {*
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   193
signature EVAL =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   194
sig
24587
4f2cbf6e563f multi-functional value keyword
haftmann
parents: 24508
diff changeset
   195
  val eval_ref: (unit -> term) option ref
24835
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   196
  val eval_term: theory -> term -> term
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   197
  val evaluate: Proof.context -> term -> unit
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   198
  val evaluate': string -> Proof.context -> term -> unit
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   199
  val evaluate_cmd: string option -> Toplevel.state -> unit
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   200
end;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   201
24280
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   202
structure Eval =
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   203
struct
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   204
24587
4f2cbf6e563f multi-functional value keyword
haftmann
parents: 24508
diff changeset
   205
val eval_ref = ref (NONE : (unit -> term) option);
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   206
24835
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   207
fun eval_invoke thy code ((_, ty), t) deps _ =
24916
dc56dd1b3cda simplified evaluation
haftmann
parents: 24867
diff changeset
   208
  CodeTarget.eval_invoke thy ("Eval.eval_ref", eval_ref) code (t, ty) [];
24280
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   209
24835
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   210
fun eval_term thy =
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   211
  TermOf.mk_term_of
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   212
  #> CodePackage.eval_term thy (eval_invoke thy)
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   213
  #> Code.postprocess_term thy;
24280
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   214
24835
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   215
val evaluators = [
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   216
  ("code", eval_term),
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   217
  ("SML", Codegen.eval_term),
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   218
  ("normal_form", Nbe.norm_term)
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   219
];
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   220
24835
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   221
fun gen_evaluate evaluators ctxt t =
24280
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   222
  let
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   223
    val thy = ProofContext.theory_of ctxt;
24835
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   224
    val (evls, evl) = split_last evaluators;
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   225
    val t' = case get_first (fn f => try (f thy) t) evls
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   226
     of SOME t' => t'
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   227
      | NONE => evl thy t;
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   228
    val ty' = Term.type_of t';
24920
2a45e400fdad generic Syntax.pretty/string_of operations;
wenzelm
parents: 24916
diff changeset
   229
    val p = Pretty.block [Pretty.quote (Syntax.pretty_term ctxt t'),
24835
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   230
      Pretty.fbrk, Pretty.str "::", Pretty.brk 1,
24920
2a45e400fdad generic Syntax.pretty/string_of operations;
wenzelm
parents: 24916
diff changeset
   231
      Pretty.quote (Syntax.pretty_typ ctxt ty')];
24280
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   232
  in Pretty.writeln p end;
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   233
24835
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   234
val evaluate = gen_evaluate (map snd evaluators);
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   235
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   236
fun evaluate' name = gen_evaluate
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   237
  [(the o AList.lookup (op =) evaluators) name];
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   238
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   239
fun evaluate_cmd some_name raw_t state =
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   240
  let
22804
haftmann
parents: 22665
diff changeset
   241
    val ctxt = Toplevel.context_of state;
24508
c8b82fec6447 replaced ProofContext.read_term/prop by general Syntax.read_term/prop;
wenzelm
parents: 24423
diff changeset
   242
    val t = Syntax.read_term ctxt raw_t;
24835
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   243
  in case some_name
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   244
   of NONE => evaluate ctxt t
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   245
    | SOME name => evaluate' name ctxt t
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   246
  end;
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   247
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   248
end;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   249
*}
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   250
22804
haftmann
parents: 22665
diff changeset
   251
ML {*
haftmann
parents: 22665
diff changeset
   252
  OuterSyntax.improper_command "value" "read, evaluate and print term" OuterKeyword.diag
24835
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   253
    (Scan.option (OuterParse.$$$ "(" |-- OuterParse.name --| OuterParse.$$$ ")")
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   254
    -- OuterParse.term
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   255
      >> (fn (some_name, t) => Toplevel.no_timing o Toplevel.keep
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   256
           (Eval.evaluate_cmd some_name t)));
22804
haftmann
parents: 22665
diff changeset
   257
*}
haftmann
parents: 22665
diff changeset
   258
22665
cf152ff55d16 tuned document (headers, sections, spacing);
wenzelm
parents: 22527
diff changeset
   259
end