src/Tools/Code/code_runtime.ML
changeset 58558 30cc7ee5ac10
parent 58557 fea97f7be494
child 58643 133203bd474b
equal deleted inserted replaced
58557:fea97f7be494 58558:30cc7ee5ac10
    26   val static_value_exn: 'a cookie -> { ctxt: Proof.context, target: string option,
    26   val static_value_exn: 'a cookie -> { ctxt: Proof.context, target: string option,
    27     lift_postproc: (term -> term) -> 'a -> 'a, consts: string list }
    27     lift_postproc: (term -> term) -> 'a -> 'a, consts: string list }
    28     -> Proof.context -> term -> 'a Exn.result
    28     -> Proof.context -> term -> 'a Exn.result
    29   val dynamic_holds_conv: Proof.context -> conv
    29   val dynamic_holds_conv: Proof.context -> conv
    30   val static_holds_conv: { ctxt: Proof.context, consts: string list } -> Proof.context -> conv
    30   val static_holds_conv: { ctxt: Proof.context, consts: string list } -> Proof.context -> conv
       
    31   val static_value_exn': (Proof.context -> term -> 'a) cookie
       
    32     -> { ctxt: Proof.context, lift_postproc: (term -> term) -> 'a -> 'a,
       
    33            consts: (string * typ) list, T: typ }
       
    34     -> Proof.context -> term -> 'a Exn.result (*EXPERIMENTAL!*)
    31   val code_reflect: (string * string list option) list -> string list -> string
    35   val code_reflect: (string * string list option) list -> string list -> string
    32     -> string option -> theory -> theory
    36     -> string option -> theory -> theory
    33   datatype truth = Holds
    37   datatype truth = Holds
    34   val put_truth: (unit -> truth) -> Proof.context -> Proof.context
    38   val put_truth: (unit -> truth) -> Proof.context -> Proof.context
    35   val trace: bool Config.T
    39   val trace: bool Config.T
   152   fun init _ () = error "Truth_Result"
   156   fun init _ () = error "Truth_Result"
   153 );
   157 );
   154 val put_truth = Truth_Result.put;
   158 val put_truth = Truth_Result.put;
   155 val truth_cookie = (Truth_Result.get, put_truth, Long_Name.append this "put_truth");
   159 val truth_cookie = (Truth_Result.get, put_truth, Long_Name.append this "put_truth");
   156 
   160 
       
   161 local
       
   162 
   157 val reject_vars = fn ctxt => tap (reject_vars ctxt o Thm.term_of);
   163 val reject_vars = fn ctxt => tap (reject_vars ctxt o Thm.term_of);
   158 
       
   159 local
       
   160 
   164 
   161 fun check_holds ctxt evaluator vs_t ct =
   165 fun check_holds ctxt evaluator vs_t ct =
   162   let
   166   let
   163     val thy = Proof_Context.theory_of ctxt;
   167     val thy = Proof_Context.theory_of ctxt;
   164     val t = Thm.term_of ct;
   168     val t = Thm.term_of ct;
   192     K (check_holds_oracle ctxt' (obtain_evaluator ctxt NONE program (map Constant deps)) vs_t o reject_vars ctxt'));
   196     K (check_holds_oracle ctxt' (obtain_evaluator ctxt NONE program (map Constant deps)) vs_t o reject_vars ctxt'));
   193 
   197 
   194 end; (*local*)
   198 end; (*local*)
   195 
   199 
   196 
   200 
   197 (** instrumentalization **)
   201 (** full static evaluation -- still with limited coverage! **)
   198 
   202 
   199 fun evaluation_code ctxt module_name program tycos consts =
   203 fun evaluation_code ctxt module_name program tycos consts =
   200   let
   204   let
   201     val thy = Proof_Context.theory_of ctxt;
   205     val thy = Proof_Context.theory_of ctxt;
   202     val (ml_modules, target_names) =
   206     val (ml_modules, target_names) =
   214           error ("Type " ^ quote (Proof_Context.markup_type ctxt tyco) ^
   218           error ("Type " ^ quote (Proof_Context.markup_type ctxt tyco) ^
   215             "\nhas a user-defined serialization")
   219             "\nhas a user-defined serialization")
   216         | SOME tyco' => (tyco, tyco')) tycos tycos';
   220         | SOME tyco' => (tyco, tyco')) tycos tycos';
   217   in (ml_code, (tycos_map, consts_map)) end;
   221   in (ml_code, (tycos_map, consts_map)) end;
   218 
   222 
   219 
   223 fun typ_signatures_for T =
   220 (* by antiquotation *)
   224   let
       
   225     val (Ts, T') = strip_type T;
       
   226   in map_range (fn n => (drop n Ts ---> T', take n Ts)) (length Ts + 1) end;
       
   227 
       
   228 fun typ_signatures cTs =
       
   229   let
       
   230     fun add (c, T) =
       
   231       fold (fn (T, Ts) => Typtab.map_default (T, []) (cons (c, Ts)))
       
   232         (typ_signatures_for T);
       
   233   in
       
   234     Typtab.empty
       
   235     |> fold add cTs
       
   236     |> Typtab.lookup_list
       
   237   end;
       
   238 
       
   239 fun print_of_term_funs { typ_sign_for, ml_name_for_const, ml_name_for_typ } Ts =
       
   240   let
       
   241     val var_names = map_range (fn n => "t" ^ string_of_int (n + 1));
       
   242     fun print_lhs c xs = "Const (" ^ quote c ^ ", _)"
       
   243       |> fold (fn x => fn s => s ^ " $ " ^ x) xs
       
   244       |> enclose "(" ")"
       
   245       |> prefix "ctxt ";
       
   246     fun print_rhs c Ts xs = ml_name_for_const c
       
   247       |> fold2 (fn T => fn x => fn s =>
       
   248          s ^ (" (" ^ ml_name_for_typ T ^ " ctxt " ^ x ^ ")")) Ts xs
       
   249     fun print_eq (c, Ts) =
       
   250       let
       
   251         val xs = var_names (length Ts);
       
   252       in print_lhs c xs ^ " = " ^ print_rhs c Ts xs end;
       
   253     val err_eq =
       
   254       "ctxt t = error (" ^ quote "Bad term: " ^ " ^ Syntax.string_of_term ctxt t)";
       
   255     fun print_eqs T =
       
   256       let
       
   257         val typ_signs = typ_sign_for T;
       
   258         val name = ml_name_for_typ T;
       
   259       in
       
   260         (map print_eq typ_signs @ [err_eq])
       
   261         |> map (prefix (name ^ " "))
       
   262         |> space_implode "\n  | "
       
   263       end;
       
   264   in
       
   265     map print_eqs Ts
       
   266     |> space_implode "\nand "
       
   267     |> prefix "fun "
       
   268     |> pair (map ml_name_for_typ Ts)
       
   269   end;
       
   270 
       
   271 fun print_of_term ctxt ml_name_for_const T cTs =
       
   272   let
       
   273     val typ_sign_for = typ_signatures cTs;
       
   274     fun add_typ T Ts =
       
   275       if member (op =) Ts T
       
   276       then Ts
       
   277       else Ts
       
   278         |> cons T
       
   279         |> fold (fold add_typ o snd) (typ_sign_for T);
       
   280     val Ts = add_typ T [];
       
   281     fun tycos_of (Type (tyco, Ts)) = maps tycos_of Ts @ [tyco]
       
   282       | tycos_of _ = [];
       
   283     val ml_name_of = Name.desymbolize NONE o Long_Name.base_name;
       
   284     val ml_names = map (suffix "_of_term" o space_implode "_" o map ml_name_of o tycos_of) Ts
       
   285       |> Name.variant_list [];
       
   286     val ml_name_for_typ = the o AList.lookup (op =) (Ts ~~ ml_names);
       
   287   in
       
   288     print_of_term_funs { typ_sign_for = typ_sign_for,
       
   289       ml_name_for_const = ml_name_for_const,
       
   290       ml_name_for_typ = ml_name_for_typ } Ts
       
   291   end;
       
   292 
       
   293 fun compile_evaluator cookie ctxt cs_code cTs T { program, deps } =
       
   294   let
       
   295     val (context_code, (_, const_map)) =
       
   296       evaluation_code ctxt structure_generated program [] cs_code;
       
   297     val ml_name_for_const = the o AList.lookup (op =) const_map;
       
   298     val (ml_names, of_term_code) = print_of_term ctxt ml_name_for_const T cTs
       
   299     val of_term = value ctxt cookie (context_code ^ "\n" ^ of_term_code, List.last ml_names);
       
   300   in fn ctxt' => fn t => fn _ => fn _ => Exn.interruptible_capture (of_term ctxt') t end;
       
   301 
       
   302 fun static_value_exn' cookie { ctxt, lift_postproc, consts, T } =
       
   303   let
       
   304     val thy = Proof_Context.theory_of ctxt;
       
   305     val cs_code = map (Axclass.unoverload_const thy) consts;
       
   306     val cTs = map2 (fn (_, T) => fn c => (c, T)) consts cs_code;
       
   307     val evaluator = Code_Thingol.static_value { ctxt = ctxt,
       
   308       lift_postproc = Exn.map_result o lift_postproc, consts = cs_code }
       
   309       (compile_evaluator cookie ctxt cs_code cTs T);
       
   310   in fn ctxt' =>
       
   311     evaluator ctxt' o reject_vars ctxt' o Syntax.check_term ctxt' o Type.constraint T
       
   312   end;
       
   313 
       
   314 
       
   315 (** code antiquotation **)
   221 
   316 
   222 local
   317 local
   223 
   318 
   224 structure Code_Antiq_Data = Proof_Data
   319 structure Code_Antiq_Data = Proof_Data
   225 (
   320 (
   261   in (print_code is_first const, register_const const ctxt) end;
   356   in (print_code is_first const, register_const const ctxt) end;
   262 
   357 
   263 end; (*local*)
   358 end; (*local*)
   264 
   359 
   265 
   360 
   266 (* reflection support *)
   361 (** reflection support **)
   267 
   362 
   268 fun check_datatype thy tyco some_consts =
   363 fun check_datatype thy tyco some_consts =
   269   let
   364   let
   270     val constrs = (map fst o snd o fst o Code.get_type thy) tyco;
   365     val constrs = (map fst o snd o fst o Code.get_type thy) tyco;
   271     val _ = case some_consts
   366     val _ = case some_consts