src/HOL/Code_Eval.thy
changeset 32753 5fae12e4679c
parent 32752 f65d74a264dd
parent 32658 82956a3f0e0b
child 32754 4e0256f7d219
equal deleted inserted replaced
32752:f65d74a264dd 32753:5fae12e4679c
     1 (*  Title:      HOL/Code_Eval.thy
       
     2     Author:     Florian Haftmann, TU Muenchen
       
     3 *)
       
     4 
       
     5 header {* Term evaluation using the generic code generator *}
       
     6 
       
     7 theory Code_Eval
       
     8 imports Plain Typerep Code_Numeral
       
     9 begin
       
    10 
       
    11 subsection {* Term representation *}
       
    12 
       
    13 subsubsection {* Terms and class @{text term_of} *}
       
    14 
       
    15 datatype "term" = dummy_term
       
    16 
       
    17 definition Const :: "String.literal \<Rightarrow> typerep \<Rightarrow> term" where
       
    18   "Const _ _ = dummy_term"
       
    19 
       
    20 definition App :: "term \<Rightarrow> term \<Rightarrow> term" where
       
    21   "App _ _ = dummy_term"
       
    22 
       
    23 code_datatype Const App
       
    24 
       
    25 class term_of = typerep +
       
    26   fixes term_of :: "'a \<Rightarrow> term"
       
    27 
       
    28 lemma term_of_anything: "term_of x \<equiv> t"
       
    29   by (rule eq_reflection) (cases "term_of x", cases t, simp)
       
    30 
       
    31 definition valapp :: "('a \<Rightarrow> 'b) \<times> (unit \<Rightarrow> term)
       
    32   \<Rightarrow> 'a \<times> (unit \<Rightarrow> term) \<Rightarrow> 'b \<times> (unit \<Rightarrow> term)" where
       
    33   "valapp f x = (fst f (fst x), \<lambda>u. App (snd f ()) (snd x ()))"
       
    34 
       
    35 lemma valapp_code [code, code_unfold]:
       
    36   "valapp (f, tf) (x, tx) = (f x, \<lambda>u. App (tf ()) (tx ()))"
       
    37   by (simp only: valapp_def fst_conv snd_conv)
       
    38 
       
    39 
       
    40 subsubsection {* @{text term_of} instances *}
       
    41 
       
    42 instantiation "fun" :: (typerep, typerep) term_of
       
    43 begin
       
    44 
       
    45 definition
       
    46   "term_of (f \<Colon> 'a \<Rightarrow> 'b) = Const (STR ''dummy_pattern'') (Typerep.Typerep (STR ''fun'')
       
    47      [Typerep.typerep TYPE('a), Typerep.typerep TYPE('b)])"
       
    48 
       
    49 instance ..
       
    50 
       
    51 end
       
    52 
       
    53 setup {*
       
    54 let
       
    55   fun add_term_of tyco raw_vs thy =
       
    56     let
       
    57       val vs = map (fn (v, _) => (v, @{sort typerep})) raw_vs;
       
    58       val ty = Type (tyco, map TFree vs);
       
    59       val lhs = Const (@{const_name term_of}, ty --> @{typ term})
       
    60         $ Free ("x", ty);
       
    61       val rhs = @{term "undefined \<Colon> term"};
       
    62       val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));
       
    63       fun triv_name_of t = (fst o dest_Free o fst o strip_comb o fst
       
    64         o HOLogic.dest_eq o HOLogic.dest_Trueprop) t ^ "_triv";
       
    65     in
       
    66       thy
       
    67       |> TheoryTarget.instantiation ([tyco], vs, @{sort term_of})
       
    68       |> `(fn lthy => Syntax.check_term lthy eq)
       
    69       |-> (fn eq => Specification.definition (NONE, ((Binding.name (triv_name_of eq), []), eq)))
       
    70       |> snd
       
    71       |> Class.prove_instantiation_exit (K (Class.intro_classes_tac []))
       
    72     end;
       
    73   fun ensure_term_of (tyco, (raw_vs, _)) thy =
       
    74     let
       
    75       val need_inst = not (can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort term_of})
       
    76         andalso can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort typerep};
       
    77     in if need_inst then add_term_of tyco raw_vs thy else thy end;
       
    78 in
       
    79   Code.type_interpretation ensure_term_of
       
    80 end
       
    81 *}
       
    82 
       
    83 setup {*
       
    84 let
       
    85   fun mk_term_of_eq thy ty vs tyco (c, tys) =
       
    86     let
       
    87       val t = list_comb (Const (c, tys ---> ty),
       
    88         map Free (Name.names Name.context "a" tys));
       
    89       val (arg, rhs) = pairself (Thm.cterm_of thy o map_types Logic.unvarifyT o Logic.varify)
       
    90         (t, (map_aterms (fn t as Free (v, ty) => HOLogic.mk_term_of ty t | t => t) o HOLogic.reflect_term) t)
       
    91       val cty = Thm.ctyp_of thy ty;
       
    92     in
       
    93       @{thm term_of_anything}
       
    94       |> Drule.instantiate' [SOME cty] [SOME arg, SOME rhs]
       
    95       |> Thm.varifyT
       
    96     end;
       
    97   fun add_term_of_code tyco raw_vs raw_cs thy =
       
    98     let
       
    99       val algebra = Sign.classes_of thy;
       
   100       val vs = map (fn (v, sort) =>
       
   101         (v, curry (Sorts.inter_sort algebra) @{sort typerep} sort)) raw_vs;
       
   102       val ty = Type (tyco, map TFree vs);
       
   103       val cs = (map o apsnd o map o map_atyps)
       
   104         (fn TFree (v, _) => TFree (v, (the o AList.lookup (op =) vs) v)) raw_cs;
       
   105       val const = AxClass.param_of_inst thy (@{const_name term_of}, tyco);
       
   106       val eqs = map (mk_term_of_eq thy ty vs tyco) cs;
       
   107    in
       
   108       thy
       
   109       |> Code.del_eqns const
       
   110       |> fold Code.add_eqn eqs
       
   111     end;
       
   112   fun ensure_term_of_code (tyco, (raw_vs, cs)) thy =
       
   113     let
       
   114       val has_inst = can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort term_of};
       
   115     in if has_inst then add_term_of_code tyco raw_vs cs thy else thy end;
       
   116 in
       
   117   Code.type_interpretation ensure_term_of_code
       
   118 end
       
   119 *}
       
   120 
       
   121 
       
   122 subsubsection {* Code generator setup *}
       
   123 
       
   124 lemmas [code del] = term.recs term.cases term.size
       
   125 lemma [code, code del]: "eq_class.eq (t1\<Colon>term) t2 \<longleftrightarrow> eq_class.eq t1 t2" ..
       
   126 
       
   127 lemma [code, code del]: "(term_of \<Colon> typerep \<Rightarrow> term) = term_of" ..
       
   128 lemma [code, code del]: "(term_of \<Colon> term \<Rightarrow> term) = term_of" ..
       
   129 lemma [code, code del]: "(term_of \<Colon> String.literal \<Rightarrow> term) = term_of" ..
       
   130 lemma [code, code del]:
       
   131   "(Code_Eval.term_of \<Colon> 'a::{type, term_of} Predicate.pred \<Rightarrow> Code_Eval.term) = Code_Eval.term_of" ..
       
   132 lemma [code, code del]:
       
   133   "(Code_Eval.term_of \<Colon> 'a::{type, term_of} Predicate.seq \<Rightarrow> Code_Eval.term) = Code_Eval.term_of" ..
       
   134 
       
   135 lemma term_of_char [unfolded typerep_fun_def typerep_char_def typerep_nibble_def, code]: "Code_Eval.term_of c =
       
   136     (let (n, m) = nibble_pair_of_char c
       
   137   in Code_Eval.App (Code_Eval.App (Code_Eval.Const (STR ''String.char.Char'') (TYPEREP(nibble \<Rightarrow> nibble \<Rightarrow> char)))
       
   138     (Code_Eval.term_of n)) (Code_Eval.term_of m))"
       
   139   by (subst term_of_anything) rule 
       
   140 
       
   141 code_type "term"
       
   142   (Eval "Term.term")
       
   143 
       
   144 code_const Const and App
       
   145   (Eval "Term.Const/ ((_), (_))" and "Term.$/ ((_), (_))")
       
   146 
       
   147 code_const "term_of \<Colon> String.literal \<Rightarrow> term"
       
   148   (Eval "HOLogic.mk'_message'_string")
       
   149 
       
   150 code_reserved Eval HOLogic
       
   151 
       
   152 
       
   153 subsubsection {* Syntax *}
       
   154 
       
   155 definition termify :: "'a \<Rightarrow> term" where
       
   156   [code del]: "termify x = dummy_term"
       
   157 
       
   158 abbreviation valtermify :: "'a \<Rightarrow> 'a \<times> (unit \<Rightarrow> term)" where
       
   159   "valtermify x \<equiv> (x, \<lambda>u. termify x)"
       
   160 
       
   161 setup {*
       
   162 let
       
   163   fun map_default f xs =
       
   164     let val ys = map f xs
       
   165     in if exists is_some ys
       
   166       then SOME (map2 the_default xs ys)
       
   167       else NONE
       
   168     end;
       
   169   fun subst_termify_app (Const (@{const_name termify}, T), [t]) =
       
   170         if not (Term.has_abs t)
       
   171         then if fold_aterms (fn Const _ => I | _ => K false) t true
       
   172           then SOME (HOLogic.reflect_term t)
       
   173           else error "Cannot termify expression containing variables"
       
   174         else error "Cannot termify expression containing abstraction"
       
   175     | subst_termify_app (t, ts) = case map_default subst_termify ts
       
   176        of SOME ts' => SOME (list_comb (t, ts'))
       
   177         | NONE => NONE
       
   178   and subst_termify (Abs (v, T, t)) = (case subst_termify t
       
   179        of SOME t' => SOME (Abs (v, T, t'))
       
   180         | NONE => NONE)
       
   181     | subst_termify t = subst_termify_app (strip_comb t) 
       
   182   fun check_termify ts ctxt = map_default subst_termify ts
       
   183     |> Option.map (rpair ctxt)
       
   184 in
       
   185   Context.theory_map (Syntax.add_term_check 0 "termify" check_termify)
       
   186 end;
       
   187 *}
       
   188 
       
   189 locale term_syntax
       
   190 begin
       
   191 
       
   192 notation App (infixl "<\<cdot>>" 70)
       
   193   and valapp (infixl "{\<cdot>}" 70)
       
   194 
       
   195 end
       
   196 
       
   197 interpretation term_syntax .
       
   198 
       
   199 no_notation App (infixl "<\<cdot>>" 70)
       
   200   and valapp (infixl "{\<cdot>}" 70)
       
   201 
       
   202 
       
   203 subsection {* Numeric types *}
       
   204 
       
   205 definition term_of_num :: "'a\<Colon>{semiring_div} \<Rightarrow> 'a\<Colon>{semiring_div} \<Rightarrow> term" where
       
   206   "term_of_num two = (\<lambda>_. dummy_term)"
       
   207 
       
   208 lemma (in term_syntax) term_of_num_code [code]:
       
   209   "term_of_num two k = (if k = 0 then termify Int.Pls
       
   210     else (if k mod two = 0
       
   211       then termify Int.Bit0 <\<cdot>> term_of_num two (k div two)
       
   212       else termify Int.Bit1 <\<cdot>> term_of_num two (k div two)))"
       
   213   by (auto simp add: term_of_anything Const_def App_def term_of_num_def Let_def)
       
   214 
       
   215 lemma (in term_syntax) term_of_nat_code [code]:
       
   216   "term_of (n::nat) = termify (number_of :: int \<Rightarrow> nat) <\<cdot>> term_of_num (2::nat) n"
       
   217   by (simp only: term_of_anything)
       
   218 
       
   219 lemma (in term_syntax) term_of_int_code [code]:
       
   220   "term_of (k::int) = (if k = 0 then termify (0 :: int)
       
   221     else if k > 0 then termify (number_of :: int \<Rightarrow> int) <\<cdot>> term_of_num (2::int) k
       
   222       else termify (uminus :: int \<Rightarrow> int) <\<cdot>> (termify (number_of :: int \<Rightarrow> int) <\<cdot>> term_of_num (2::int) (- k)))"
       
   223   by (simp only: term_of_anything)
       
   224 
       
   225 lemma (in term_syntax) term_of_code_numeral_code [code]:
       
   226   "term_of (k::code_numeral) = termify (number_of :: int \<Rightarrow> code_numeral) <\<cdot>> term_of_num (2::code_numeral) k"
       
   227   by (simp only: term_of_anything)
       
   228 
       
   229 subsection {* Obfuscate *}
       
   230 
       
   231 print_translation {*
       
   232 let
       
   233   val term = Const ("<TERM>", dummyT);
       
   234   fun tr1' [_, _] = term;
       
   235   fun tr2' [] = term;
       
   236 in
       
   237   [(@{const_syntax Const}, tr1'),
       
   238     (@{const_syntax App}, tr1'),
       
   239     (@{const_syntax dummy_term}, tr2')]
       
   240 end
       
   241 *}
       
   242 
       
   243 hide const dummy_term App valapp
       
   244 hide (open) const Const termify valtermify term_of term_of_num
       
   245 
       
   246 
       
   247 subsection {* Evaluation setup *}
       
   248 
       
   249 ML {*
       
   250 signature EVAL =
       
   251 sig
       
   252   val eval_ref: (unit -> term) option ref
       
   253   val eval_term: theory -> term -> term
       
   254 end;
       
   255 
       
   256 structure Eval : EVAL =
       
   257 struct
       
   258 
       
   259 val eval_ref = ref (NONE : (unit -> term) option);
       
   260 
       
   261 fun eval_term thy t =
       
   262   Code_ML.eval NONE ("Eval.eval_ref", eval_ref) I thy (HOLogic.mk_term_of (fastype_of t) t) [];
       
   263 
       
   264 end;
       
   265 *}
       
   266 
       
   267 setup {*
       
   268   Value.add_evaluator ("code", Eval.eval_term o ProofContext.theory_of)
       
   269 *}
       
   270 
       
   271 end