src/HOL/Library/Eval.thy
author haftmann
Thu, 04 Oct 2007 19:41:49 +0200
changeset 24835 8c26128f8997
parent 24659 6b7ac2a43df8
child 24867 e5b55d7be9bb
permissions -rw-r--r--
clarified relationship of code generator conversions and evaluations
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
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
24587
4f2cbf6e563f multi-functional value keyword
haftmann
parents: 24508
diff changeset
    37
instance "prop" :: typ_of
4f2cbf6e563f multi-functional value keyword
haftmann
parents: 24508
diff changeset
    38
  "typ_of T \<equiv> STR ''prop'' {\<struct>} []" ..
4f2cbf6e563f multi-functional value keyword
haftmann
parents: 24508
diff changeset
    39
23062
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
    40
instance itself :: (typ_of) typ_of
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
    41
  "typ_of T \<equiv> STR ''itself'' {\<struct>} [typ_of TYPE('a\<Colon>typ_of)]" ..
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
    42
24587
4f2cbf6e563f multi-functional value keyword
haftmann
parents: 24508
diff changeset
    43
instance set :: (typ_of) typ_of
4f2cbf6e563f multi-functional value keyword
haftmann
parents: 24508
diff changeset
    44
  "typ_of T \<equiv> STR ''set'' {\<struct>} [typ_of TYPE('a\<Colon>typ_of)]" ..
23062
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
    45
23020
abecb6a8cea6 typ_of instance for int
haftmann
parents: 22845
diff changeset
    46
instance int :: typ_of
abecb6a8cea6 typ_of instance for int
haftmann
parents: 22845
diff changeset
    47
  "typ_of T \<equiv> STR ''IntDef.int'' {\<struct>} []" ..
abecb6a8cea6 typ_of instance for int
haftmann
parents: 22845
diff changeset
    48
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    49
setup {*
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    50
let
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    51
  fun mk arities _ thy =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    52
    (maps (fn (tyco, asorts, _) => [(("", []), TypOf.mk_typ_of_def
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    53
      (Type (tyco,
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    54
        map TFree (Name.names Name.context "'a" asorts))))]) arities, thy);
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    55
  fun hook specs =
24219
e558fe311376 new structure for code generator modules
haftmann
parents: 23133
diff changeset
    56
    DatatypeCodegen.prove_codetypes_arities (Class.intro_classes_tac [])
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    57
      (map (fn (tyco, (is_dt, _)) => (tyco, is_dt)) specs)
24621
97d403d9ab54 clarified evaluation code
haftmann
parents: 24587
diff changeset
    58
      [TypOf.class_typ_of] mk ((K o K) (fold Code.add_default_func))
24626
85eceef2edc7 introduced generic concepts for theory interpretators
haftmann
parents: 24621
diff changeset
    59
in DatatypeCodegen.add_codetypes_hook hook end
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    60
*}
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    61
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    62
22527
84690fcd3db9 fixed document preparation
haftmann
parents: 22525
diff changeset
    63
subsection {* @{text term_of} class *}
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    64
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    65
class term_of = typ_of +
23062
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
    66
  constrains typ_of :: "'a\<Colon>{} itself \<Rightarrow> typ"
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    67
  fixes term_of :: "'a \<Rightarrow> term"
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    68
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    69
ML {*
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    70
structure TermOf =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    71
struct
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    72
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    73
local
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    74
  fun term_term_of ty =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    75
    Const (@{const_name term_of}, ty --> @{typ term});
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    76
in
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    77
  val class_term_of = Sign.intern_class @{theory} "term_of";
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    78
  fun mk_terms_of_defs vs (tyco, cs) =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    79
    let
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    80
      val dty = Type (tyco, map TFree vs);
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    81
      fun mk_eq c =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    82
        let
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    83
          val lhs : term = term_term_of dty $ c;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    84
          val rhs : term = Pure_term.mk_term
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    85
            (fn (v, ty) => term_term_of ty $ Free (v, ty))
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    86
            (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
    87
        in
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    88
          HOLogic.mk_eq (lhs, rhs)
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    89
        end;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    90
    in map mk_eq cs end;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    91
  fun mk_term_of t =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    92
    term_term_of (Term.fastype_of t) $ t;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    93
end;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    94
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    95
end;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    96
*}
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    97
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    98
setup {*
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    99
let
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   100
  fun thy_note ((name, atts), thms) =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   101
    PureThy.add_thmss [((name, thms), atts)] #-> (fn [thms] => pair (name, thms));
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   102
  fun thy_def ((name, atts), t) =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   103
    PureThy.add_defs_i false [((name, t), atts)] #-> (fn [thm] => pair (name, thm));
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24381
diff changeset
   104
  fun mk arities css _ thy =
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   105
    let
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   106
      val (_, asorts, _) :: _ = arities;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   107
      val vs = Name.names Name.context "'a" asorts;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   108
      val defs = map (TermOf.mk_terms_of_defs vs) css;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   109
      val defs' = (map (pair ("", []) o ObjectLogic.ensure_propT thy) o flat) defs;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   110
    in
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   111
      thy
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   112
      |> PrimrecPackage.gen_primrec thy_note thy_def "" defs'
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   113
      |> snd
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   114
    end;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   115
  fun hook specs =
24659
6b7ac2a43df8 more permissive
haftmann
parents: 24626
diff changeset
   116
    if null specs orelse (fst o hd) specs = (fst o dest_Type) @{typ typ} then I
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   117
    else
24219
e558fe311376 new structure for code generator modules
haftmann
parents: 23133
diff changeset
   118
      DatatypeCodegen.prove_codetypes_arities (Class.intro_classes_tac [])
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   119
      (map (fn (tyco, (is_dt, _)) => (tyco, is_dt)) specs)
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   120
      [TermOf.class_term_of] ((K o K o pair) []) mk
24626
85eceef2edc7 introduced generic concepts for theory interpretators
haftmann
parents: 24621
diff changeset
   121
in DatatypeCodegen.add_codetypes_hook hook end
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   122
*}
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   123
23062
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   124
abbreviation
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   125
  intT :: "typ"
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   126
where
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   127
  "intT \<equiv> STR ''IntDef.int'' {\<struct>} []"
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   128
23133
5a6935d598c3 fixed typo
haftmann
parents: 23062
diff changeset
   129
abbreviation
5a6935d598c3 fixed typo
haftmann
parents: 23062
diff changeset
   130
  bitT :: "typ"
5a6935d598c3 fixed typo
haftmann
parents: 23062
diff changeset
   131
where
5a6935d598c3 fixed typo
haftmann
parents: 23062
diff changeset
   132
  "bitT \<equiv> STR ''Numeral.bit'' {\<struct>} []"
5a6935d598c3 fixed typo
haftmann
parents: 23062
diff changeset
   133
23062
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   134
function
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   135
  mk_int :: "int \<Rightarrow> term"
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   136
where
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   137
  "mk_int k = (if k = 0 then STR ''Numeral.Pls'' \<Colon>\<subseteq> intT
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   138
    else if k = -1 then STR ''Numeral.Min'' \<Colon>\<subseteq> intT
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   139
    else let (l, m) = divAlg (k, 2)
23133
5a6935d598c3 fixed typo
haftmann
parents: 23062
diff changeset
   140
  in STR ''Numeral.Bit'' \<Colon>\<subseteq> intT \<rightarrow> bitT \<rightarrow> intT \<bullet> mk_int l \<bullet>
5a6935d598c3 fixed typo
haftmann
parents: 23062
diff changeset
   141
    (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
   142
by pat_completeness auto
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   143
termination by (relation "measure (nat o abs)") (auto simp add: divAlg_mod_div)
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   144
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   145
instance int :: term_of
23133
5a6935d598c3 fixed typo
haftmann
parents: 23062
diff changeset
   146
  "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
   147
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   148
22804
haftmann
parents: 22665
diff changeset
   149
text {* Adaption for @{typ ml_string}s *}
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   150
22845
5f9138bcb3d7 changed code generator invocation syntax
haftmann
parents: 22804
diff changeset
   151
lemmas [code func, code func del] = term_of_ml_string_def
22525
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
subsection {* Evaluation infrastructure *}
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   155
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   156
ML {*
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   157
signature EVAL =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   158
sig
24587
4f2cbf6e563f multi-functional value keyword
haftmann
parents: 24508
diff changeset
   159
  val eval_ref: (unit -> term) option ref
24835
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   160
  val eval_term: theory -> term -> term
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   161
  val evaluate: Proof.context -> term -> unit
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   162
  val evaluate': string -> Proof.context -> term -> unit
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   163
  val evaluate_cmd: string option -> Toplevel.state -> unit
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   164
end;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   165
24280
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   166
structure Eval =
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   167
struct
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   168
24587
4f2cbf6e563f multi-functional value keyword
haftmann
parents: 24508
diff changeset
   169
val eval_ref = ref (NONE : (unit -> term) option);
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   170
24835
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   171
fun eval_invoke thy code ((_, ty), t) deps _ =
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   172
  CodePackage.eval_invoke thy ("Eval.eval_ref", eval_ref) code (t, ty) [];
24280
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   173
24835
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   174
fun eval_term thy =
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   175
  TermOf.mk_term_of
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   176
  #> CodePackage.eval_term thy (eval_invoke thy)
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   177
  #> Code.postprocess_term thy;
24280
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   178
24835
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   179
val evaluators = [
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   180
  ("code", eval_term),
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   181
  ("SML", Codegen.eval_term),
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   182
  ("normal_form", Nbe.norm_term)
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   183
];
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   184
24835
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   185
fun gen_evaluate evaluators ctxt t =
24280
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   186
  let
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   187
    val thy = ProofContext.theory_of ctxt;
24835
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   188
    val (evls, evl) = split_last evaluators;
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   189
    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
   190
     of SOME t' => t'
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   191
      | NONE => evl thy t;
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   192
    val ty' = Term.type_of t';
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   193
    val p = Pretty.block [Pretty.quote (ProofContext.pretty_term ctxt t'),
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   194
      Pretty.fbrk, Pretty.str "::", Pretty.brk 1,
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   195
      Pretty.quote (ProofContext.pretty_typ ctxt ty')];
24280
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   196
  in Pretty.writeln p end;
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   197
24835
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   198
val evaluate = gen_evaluate (map snd evaluators);
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   199
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   200
fun evaluate' name = gen_evaluate
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   201
  [(the o AList.lookup (op =) evaluators) name];
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   202
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   203
fun evaluate_cmd some_name raw_t state =
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   204
  let
22804
haftmann
parents: 22665
diff changeset
   205
    val ctxt = Toplevel.context_of state;
24508
c8b82fec6447 replaced ProofContext.read_term/prop by general Syntax.read_term/prop;
wenzelm
parents: 24423
diff changeset
   206
    val t = Syntax.read_term ctxt raw_t;
24835
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   207
  in case some_name
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   208
   of NONE => evaluate ctxt t
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   209
    | SOME name => evaluate' name ctxt t
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   210
  end;
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   211
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   212
end;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   213
*}
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   214
22804
haftmann
parents: 22665
diff changeset
   215
ML {*
haftmann
parents: 22665
diff changeset
   216
val valueP =
haftmann
parents: 22665
diff changeset
   217
  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
   218
    (Scan.option (OuterParse.$$$ "(" |-- OuterParse.name --| OuterParse.$$$ ")")
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   219
    -- OuterParse.term
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   220
      >> (fn (some_name, t) => Toplevel.no_timing o Toplevel.keep
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
   221
           (Eval.evaluate_cmd some_name t)));
22804
haftmann
parents: 22665
diff changeset
   222
haftmann
parents: 22665
diff changeset
   223
val _ = OuterSyntax.add_parsers [valueP];
haftmann
parents: 22665
diff changeset
   224
*}
haftmann
parents: 22665
diff changeset
   225
22665
cf152ff55d16 tuned document (headers, sections, spacing);
wenzelm
parents: 22527
diff changeset
   226
end