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