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