src/HOL/Library/Eval.thy
author nipkow
Sat, 01 Sep 2007 01:21:48 +0200
changeset 24506 020db6ec334a
parent 24423 ae9cd0e92423
child 24508 c8b82fec6447
permissions -rw-r--r--
final(?) iteration of sgn saga.
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
23062
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
    37
instance itself :: (typ_of) typ_of
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
    38
  "typ_of T \<equiv> STR ''itself'' {\<struct>} [typ_of TYPE('a\<Colon>typ_of)]" ..
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
    39
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
    40
instance "prop" :: typ_of
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
    41
  "typ_of T \<equiv> STR ''prop'' {\<struct>} []" ..
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
    42
23020
abecb6a8cea6 typ_of instance for int
haftmann
parents: 22845
diff changeset
    43
instance int :: typ_of
abecb6a8cea6 typ_of instance for int
haftmann
parents: 22845
diff changeset
    44
  "typ_of T \<equiv> STR ''IntDef.int'' {\<struct>} []" ..
abecb6a8cea6 typ_of instance for int
haftmann
parents: 22845
diff changeset
    45
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    46
setup {*
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    47
let
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    48
  fun mk arities _ thy =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    49
    (maps (fn (tyco, asorts, _) => [(("", []), TypOf.mk_typ_of_def
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    50
      (Type (tyco,
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    51
        map TFree (Name.names Name.context "'a" asorts))))]) arities, thy);
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    52
  fun hook specs =
24219
e558fe311376 new structure for code generator modules
haftmann
parents: 23133
diff changeset
    53
    DatatypeCodegen.prove_codetypes_arities (Class.intro_classes_tac [])
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    54
      (map (fn (tyco, (is_dt, _)) => (tyco, is_dt)) specs)
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24381
diff changeset
    55
      [TypOf.class_typ_of] mk ((K o K) (fold (Code.add_func true)))
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    56
in DatatypeCodegen.add_codetypes_hook_bootstrap hook end
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    57
*}
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    58
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    59
22527
84690fcd3db9 fixed document preparation
haftmann
parents: 22525
diff changeset
    60
subsection {* @{text term_of} class *}
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    61
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    62
class term_of = typ_of +
23062
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
    63
  constrains typ_of :: "'a\<Colon>{} itself \<Rightarrow> typ"
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    64
  fixes term_of :: "'a \<Rightarrow> term"
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    65
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    66
ML {*
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    67
structure TermOf =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    68
struct
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    69
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    70
local
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    71
  fun term_term_of ty =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    72
    Const (@{const_name term_of}, ty --> @{typ term});
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    73
in
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    74
  val class_term_of = Sign.intern_class @{theory} "term_of";
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    75
  fun mk_terms_of_defs vs (tyco, cs) =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    76
    let
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    77
      val dty = Type (tyco, map TFree vs);
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    78
      fun mk_eq c =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    79
        let
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    80
          val lhs : term = term_term_of dty $ c;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    81
          val rhs : term = Pure_term.mk_term
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    82
            (fn (v, ty) => term_term_of ty $ Free (v, ty))
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    83
            (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
    84
        in
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    85
          HOLogic.mk_eq (lhs, rhs)
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    86
        end;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    87
    in map mk_eq cs end;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    88
  fun mk_term_of t =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    89
    term_term_of (Term.fastype_of t) $ t;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    90
end;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    91
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    92
end;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    93
*}
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    94
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    95
setup {*
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    96
let
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    97
  fun thy_note ((name, atts), thms) =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    98
    PureThy.add_thmss [((name, thms), atts)] #-> (fn [thms] => pair (name, thms));
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
    99
  fun thy_def ((name, atts), t) =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   100
    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
   101
  fun mk arities css _ thy =
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   102
    let
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   103
      val (_, asorts, _) :: _ = arities;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   104
      val vs = Name.names Name.context "'a" asorts;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   105
      val defs = map (TermOf.mk_terms_of_defs vs) css;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   106
      val defs' = (map (pair ("", []) o ObjectLogic.ensure_propT thy) o flat) defs;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   107
    in
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   108
      thy
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   109
      |> PrimrecPackage.gen_primrec thy_note thy_def "" defs'
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   110
      |> snd
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   111
    end;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   112
  fun hook specs =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   113
    if (fst o hd) specs = (fst o dest_Type) @{typ typ} then I
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   114
    else
24219
e558fe311376 new structure for code generator modules
haftmann
parents: 23133
diff changeset
   115
      DatatypeCodegen.prove_codetypes_arities (Class.intro_classes_tac [])
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   116
      (map (fn (tyco, (is_dt, _)) => (tyco, is_dt)) specs)
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   117
      [TermOf.class_term_of] ((K o K o pair) []) mk
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   118
in DatatypeCodegen.add_codetypes_hook_bootstrap hook end
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   119
*}
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   120
23062
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   121
abbreviation
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   122
  intT :: "typ"
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   123
where
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   124
  "intT \<equiv> STR ''IntDef.int'' {\<struct>} []"
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   125
23133
5a6935d598c3 fixed typo
haftmann
parents: 23062
diff changeset
   126
abbreviation
5a6935d598c3 fixed typo
haftmann
parents: 23062
diff changeset
   127
  bitT :: "typ"
5a6935d598c3 fixed typo
haftmann
parents: 23062
diff changeset
   128
where
5a6935d598c3 fixed typo
haftmann
parents: 23062
diff changeset
   129
  "bitT \<equiv> STR ''Numeral.bit'' {\<struct>} []"
5a6935d598c3 fixed typo
haftmann
parents: 23062
diff changeset
   130
23062
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   131
function
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   132
  mk_int :: "int \<Rightarrow> term"
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   133
where
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   134
  "mk_int k = (if k = 0 then STR ''Numeral.Pls'' \<Colon>\<subseteq> intT
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   135
    else if k = -1 then STR ''Numeral.Min'' \<Colon>\<subseteq> intT
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   136
    else let (l, m) = divAlg (k, 2)
23133
5a6935d598c3 fixed typo
haftmann
parents: 23062
diff changeset
   137
  in STR ''Numeral.Bit'' \<Colon>\<subseteq> intT \<rightarrow> bitT \<rightarrow> intT \<bullet> mk_int l \<bullet>
5a6935d598c3 fixed typo
haftmann
parents: 23062
diff changeset
   138
    (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
   139
by pat_completeness auto
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   140
termination by (relation "measure (nat o abs)") (auto simp add: divAlg_mod_div)
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   141
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   142
instance int :: term_of
23133
5a6935d598c3 fixed typo
haftmann
parents: 23062
diff changeset
   143
  "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
   144
d88d2087436d evaluation for integers
haftmann
parents: 23020
diff changeset
   145
22804
haftmann
parents: 22665
diff changeset
   146
text {* Adaption for @{typ ml_string}s *}
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   147
22845
5f9138bcb3d7 changed code generator invocation syntax
haftmann
parents: 22804
diff changeset
   148
lemmas [code func, code func del] = term_of_ml_string_def
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   149
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   150
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   151
subsection {* Evaluation infrastructure *}
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   152
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   153
ML {*
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   154
signature EVAL =
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   155
sig
22804
haftmann
parents: 22665
diff changeset
   156
  val eval_ref: term option ref
24280
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   157
  val eval_conv: cterm -> thm
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   158
  val eval_print: (cterm -> thm) -> Proof.context -> term -> unit
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   159
  val eval_print_cmd: (cterm -> thm) -> string -> Toplevel.state -> unit
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   160
end;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   161
24280
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   162
structure Eval =
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   163
struct
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   164
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   165
val eval_ref = ref (NONE : term option);
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   166
24280
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   167
end;
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   168
*}
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   169
24381
560e8ecdf633 improved evaluation interface
haftmann
parents: 24280
diff changeset
   170
oracle eval_oracle ("term * CodeThingol.code * (CodeThingol.typscheme * CodeThingol.iterm) * cterm") =
560e8ecdf633 improved evaluation interface
haftmann
parents: 24280
diff changeset
   171
{* fn thy => fn (t0, code, ((vs, ty), t), ct) => 
24280
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   172
let
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   173
  val _ = (Term.map_types o Term.map_atyps) (fn _ =>
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   174
    error ("Term " ^ Sign.string_of_term thy t0 ^ " contains polymorphic type"))
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   175
    t0;
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   176
in Logic.mk_equals (t0, CodeTarget.eval_term thy ("Eval.eval_ref", Eval.eval_ref) code (t, ty) []) end;
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   177
*}
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   178
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   179
ML {*
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   180
structure Eval : EVAL =
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   181
struct
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   182
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   183
open Eval;
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   184
24381
560e8ecdf633 improved evaluation interface
haftmann
parents: 24280
diff changeset
   185
fun eval_invoke thy t0 code vs_ty_t _ ct = eval_oracle thy (t0, code, vs_ty_t, ct);
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   186
24280
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   187
fun eval_conv ct =
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   188
  let
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   189
    val thy = Thm.theory_of_cterm ct;
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   190
    val ct' = (Thm.cterm_of thy o TermOf.mk_term_of o Thm.term_of) ct;
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   191
  in
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   192
    CodePackage.eval_term thy
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   193
      (eval_invoke thy (Thm.term_of ct)) ct'
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   194
  end;
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   195
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   196
fun eval_print conv ctxt t =
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   197
  let
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   198
    val thy = ProofContext.theory_of ctxt;
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   199
    val ct = Thm.cterm_of thy t;
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   200
    val (_, t') = (Logic.dest_equals o Thm.prop_of o conv) ct;
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   201
    val ty = Term.type_of t';
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   202
    val p = 
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   203
      Pretty.block [Pretty.quote (ProofContext.pretty_term ctxt t'), Pretty.fbrk,
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   204
        Pretty.str "::", Pretty.brk 1, Pretty.quote (ProofContext.pretty_typ ctxt ty)];
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   205
  in Pretty.writeln p end;
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   206
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   207
fun eval_print_cmd conv raw_t state =
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   208
  let
22804
haftmann
parents: 22665
diff changeset
   209
    val ctxt = Toplevel.context_of state;
24280
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   210
    val t = ProofContext.read_term ctxt raw_t;
22804
haftmann
parents: 22665
diff changeset
   211
    val thy = ProofContext.theory_of ctxt;
24280
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   212
    val ct = Thm.cterm_of thy t;
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   213
    val (_, t') = (Logic.dest_equals o Thm.prop_of o conv) ct;
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   214
    val ty = Term.type_of t';
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   215
    val p = 
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   216
      Pretty.block [Pretty.quote (ProofContext.pretty_term ctxt t'), Pretty.fbrk,
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   217
        Pretty.str "::", Pretty.brk 1, Pretty.quote (ProofContext.pretty_typ ctxt ty)];
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   218
  in Pretty.writeln p end;
22525
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   219
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   220
end;
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   221
*}
d929a900584c cleaned up Library/ and ex/
haftmann
parents:
diff changeset
   222
22804
haftmann
parents: 22665
diff changeset
   223
ML {*
haftmann
parents: 22665
diff changeset
   224
val valueP =
haftmann
parents: 22665
diff changeset
   225
  OuterSyntax.improper_command "value" "read, evaluate and print term" OuterKeyword.diag
haftmann
parents: 22665
diff changeset
   226
    ((OuterParse.opt_keyword "overloaded" -- OuterParse.term)
24280
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   227
      >> (fn (b, t) => Toplevel.no_timing o Toplevel.keep
c9867bdf2424 updated code generator setup
haftmann
parents: 24219
diff changeset
   228
           (Eval.eval_print_cmd (if b then Eval.eval_conv else Codegen.evaluation_conv) t)));
22804
haftmann
parents: 22665
diff changeset
   229
haftmann
parents: 22665
diff changeset
   230
val _ = OuterSyntax.add_parsers [valueP];
haftmann
parents: 22665
diff changeset
   231
*}
haftmann
parents: 22665
diff changeset
   232
22665
cf152ff55d16 tuned document (headers, sections, spacing);
wenzelm
parents: 22527
diff changeset
   233
end