src/Tools/nbe.ML
author wenzelm
Sat, 16 Apr 2011 16:15:37 +0200
changeset 43232 23f352990944
parent 41720 f6ab14e61604
child 43286 13ecdb3057d8
permissions -rw-r--r--
modernized structure Proof_Context;
haftmann@24590
     1
(*  Title:      Tools/nbe.ML
haftmann@24155
     2
    Authors:    Klaus Aehlig, LMU Muenchen; Tobias Nipkow, Florian Haftmann, TU Muenchen
haftmann@24155
     3
haftmann@24839
     4
Normalization by evaluation, based on generic code generator.
haftmann@24155
     5
*)
haftmann@24155
     6
haftmann@24155
     7
signature NBE =
haftmann@24155
     8
sig
haftmann@41488
     9
  val dynamic_conv: theory -> conv
haftmann@41432
    10
  val dynamic_value: theory -> term -> term
haftmann@41432
    11
  val static_conv: theory -> string list -> conv
haftmann@41432
    12
  val static_value: theory -> string list -> term -> term
haftmann@25101
    13
wenzelm@25204
    14
  datatype Univ =
haftmann@26064
    15
      Const of int * Univ list               (*named (uninterpreted) constants*)
haftmann@25924
    16
    | DFree of string * int                  (*free (uninterpreted) dictionary parameters*)
haftmann@24155
    17
    | BVar of int * Univ list
haftmann@28350
    18
    | Abs of (int * (Univ list -> Univ)) * Univ list
haftmann@25924
    19
  val apps: Univ -> Univ list -> Univ        (*explicit applications*)
haftmann@25944
    20
  val abss: int -> (Univ list -> Univ) -> Univ
haftmann@28337
    21
                                             (*abstractions as closures*)
haftmann@41285
    22
  val same: Univ * Univ -> bool
haftmann@24155
    23
haftmann@39634
    24
  val put_result: (unit -> Univ list -> Univ list) -> Proof.context -> Proof.context
wenzelm@32740
    25
  val trace: bool Unsynchronized.ref
haftmann@25924
    26
haftmann@26747
    27
  val setup: theory -> theory
haftmann@32544
    28
  val add_const_alias: thm -> theory -> theory
haftmann@24155
    29
end;
haftmann@24155
    30
haftmann@24155
    31
structure Nbe: NBE =
haftmann@24155
    32
struct
haftmann@24155
    33
haftmann@24155
    34
(* generic non-sense *)
haftmann@24155
    35
wenzelm@32740
    36
val trace = Unsynchronized.ref false;
haftmann@32544
    37
fun traced f x = if !trace then (tracing (f x); x) else x;
haftmann@24155
    38
haftmann@24155
    39
haftmann@32544
    40
(** certificates and oracle for "trivial type classes" **)
haftmann@32544
    41
wenzelm@33522
    42
structure Triv_Class_Data = Theory_Data
haftmann@32544
    43
(
haftmann@32544
    44
  type T = (class * thm) list;
haftmann@32544
    45
  val empty = [];
haftmann@32544
    46
  val extend = I;
wenzelm@33522
    47
  fun merge data : T = AList.merge (op =) (K true) data;
haftmann@32544
    48
);
haftmann@32544
    49
haftmann@32544
    50
fun add_const_alias thm thy =
haftmann@32544
    51
  let
haftmann@32544
    52
    val (ofclass, eqn) = case try Logic.dest_equals (Thm.prop_of thm)
haftmann@32544
    53
     of SOME ofclass_eq => ofclass_eq
haftmann@32544
    54
      | _ => error ("Bad certificate: " ^ Display.string_of_thm_global thy thm);
haftmann@32544
    55
    val (T, class) = case try Logic.dest_of_class ofclass
haftmann@32544
    56
     of SOME T_class => T_class
haftmann@32544
    57
      | _ => error ("Bad certificate: " ^ Display.string_of_thm_global thy thm);
haftmann@32544
    58
    val tvar = case try Term.dest_TVar T
haftmann@32544
    59
     of SOME (tvar as (_, sort)) => if null (filter (can (AxClass.get_info thy)) sort)
haftmann@32544
    60
          then tvar
haftmann@32544
    61
          else error ("Bad sort: " ^ Display.string_of_thm_global thy thm)
haftmann@32544
    62
      | _ => error ("Bad type: " ^ Display.string_of_thm_global thy thm);
haftmann@32544
    63
    val _ = if Term.add_tvars eqn [] = [tvar] then ()
haftmann@32544
    64
      else error ("Inconsistent type: " ^ Display.string_of_thm_global thy thm);
haftmann@32544
    65
    val lhs_rhs = case try Logic.dest_equals eqn
haftmann@32544
    66
     of SOME lhs_rhs => lhs_rhs
haftmann@32544
    67
      | _ => error ("Not an equation: " ^ Syntax.string_of_term_global thy eqn);
haftmann@40608
    68
    val c_c' = case try (pairself (AxClass.unoverload_const thy o dest_Const)) lhs_rhs
haftmann@32544
    69
     of SOME c_c' => c_c'
haftmann@32544
    70
      | _ => error ("Not an equation with two constants: "
haftmann@32544
    71
          ^ Syntax.string_of_term_global thy eqn);
haftmann@32544
    72
    val _ = if the_list (AxClass.class_of_param thy (snd c_c')) = [class] then ()
haftmann@32544
    73
      else error ("Inconsistent class: " ^ Display.string_of_thm_global thy thm);
haftmann@32544
    74
  in Triv_Class_Data.map (AList.update (op =) (class, thm)) thy end;
haftmann@32544
    75
haftmann@32544
    76
local
haftmann@32544
    77
haftmann@32544
    78
val get_triv_classes = map fst o Triv_Class_Data.get;
haftmann@32544
    79
haftmann@32544
    80
val (_, triv_of_class) = Context.>>> (Context.map_theory_result
haftmann@36982
    81
  (Thm.add_oracle (Binding.name "triv_of_class", fn (thy, T, class) =>
haftmann@32544
    82
    Thm.cterm_of thy (Logic.mk_of_class (T, class)))));
haftmann@32544
    83
haftmann@32544
    84
in
haftmann@32544
    85
haftmann@32544
    86
fun lift_triv_classes_conv thy conv ct =
haftmann@32544
    87
  let
haftmann@32544
    88
    val algebra = Sign.classes_of thy;
haftmann@36982
    89
    val certT = Thm.ctyp_of thy;
haftmann@32544
    90
    val triv_classes = get_triv_classes thy;
haftmann@36982
    91
    fun additional_classes sort = filter_out (fn class => Sorts.sort_le algebra (sort, [class])) triv_classes;
haftmann@36982
    92
    fun mk_entry (v, sort) =
haftmann@36982
    93
      let
haftmann@36982
    94
        val T = TFree (v, sort);
haftmann@36982
    95
        val cT = certT T;
haftmann@36982
    96
        val triv_sort = additional_classes sort;
haftmann@36982
    97
      in
haftmann@36982
    98
        (v, (Sorts.inter_sort algebra (sort, triv_sort),
haftmann@36982
    99
          (cT, AList.make (fn class => Thm.of_class (cT, class)) sort
haftmann@36982
   100
            @ AList.make (fn class => triv_of_class (thy, T, class)) triv_sort)))
haftmann@36982
   101
      end;
haftmann@36982
   102
    val vs_tab = map mk_entry (Term.add_tfrees (Thm.term_of ct) []);
haftmann@36982
   103
    fun instantiate thm =
haftmann@36982
   104
      let
haftmann@36982
   105
        val cert_tvars = map (certT o TVar) (Term.add_tvars
haftmann@36982
   106
          ((fst o Logic.dest_equals o Logic.strip_imp_concl o Thm.prop_of) thm) []);
haftmann@36982
   107
        val instantiation =
haftmann@36982
   108
          map2 (fn cert_tvar => fn (_, (_, (cT, _))) => (cert_tvar, cT)) cert_tvars vs_tab;
haftmann@36982
   109
      in Thm.instantiate (instantiation, []) thm end;
haftmann@36982
   110
    fun of_class (TFree (v, _), class) =
haftmann@36982
   111
          the (AList.lookup (op =) ((snd o snd o the o AList.lookup (op =) vs_tab) v) class)
haftmann@36982
   112
      | of_class (T, _) = error ("Bad type " ^ Syntax.string_of_typ_global thy T);
haftmann@32544
   113
    fun strip_of_class thm =
haftmann@32544
   114
      let
haftmann@36982
   115
        val prems_of_class = Thm.prop_of thm
haftmann@36982
   116
          |> Logic.strip_imp_prems
haftmann@36982
   117
          |> map (Logic.dest_of_class #> of_class);
haftmann@36982
   118
      in fold Thm.elim_implies prems_of_class thm end;
wenzelm@36779
   119
  in
wenzelm@36779
   120
    ct
haftmann@36982
   121
    |> (Drule.cterm_fun o map_types o map_type_tfree)
haftmann@36982
   122
        (fn (v, sort) => TFree (v, (fst o the o AList.lookup (op =) vs_tab) v))
haftmann@32544
   123
    |> conv
haftmann@36982
   124
    |> Thm.strip_shyps
wenzelm@35845
   125
    |> Thm.varifyT_global
haftmann@36982
   126
    |> Thm.unconstrainT
haftmann@36982
   127
    |> instantiate
haftmann@32544
   128
    |> strip_of_class
haftmann@32544
   129
  end;
haftmann@32544
   130
haftmann@32544
   131
fun lift_triv_classes_rew thy rew t =
haftmann@32544
   132
  let
haftmann@32544
   133
    val algebra = Sign.classes_of thy;
haftmann@32544
   134
    val triv_classes = get_triv_classes thy;
haftmann@32544
   135
    val vs = Term.add_tfrees t [];
haftmann@32544
   136
  in t
haftmann@32544
   137
    |> (map_types o map_type_tfree)
haftmann@32544
   138
        (fn (v, sort) => TFree (v, Sorts.inter_sort algebra (sort, triv_classes)))
haftmann@32544
   139
    |> rew
haftmann@32544
   140
    |> (map_types o map_type_tfree)
haftmann@32544
   141
        (fn (v, _) => TFree (v, the (AList.lookup (op =) vs v)))
haftmann@32544
   142
  end;
haftmann@32544
   143
haftmann@32544
   144
end;
haftmann@32544
   145
haftmann@32544
   146
haftmann@32544
   147
(** the semantic universe **)
haftmann@24155
   148
haftmann@24155
   149
(*
haftmann@24155
   150
   Functions are given by their semantical function value. To avoid
haftmann@24155
   151
   trouble with the ML-type system, these functions have the most
haftmann@24155
   152
   generic type, that is "Univ list -> Univ". The calling convention is
haftmann@24155
   153
   that the arguments come as a list, the last argument first. In
haftmann@24155
   154
   other words, a function call that usually would look like
haftmann@24155
   155
haftmann@24155
   156
   f x_1 x_2 ... x_n   or   f(x_1,x_2, ..., x_n)
haftmann@24155
   157
haftmann@24155
   158
   would be in our convention called as
haftmann@24155
   159
haftmann@24155
   160
              f [x_n,..,x_2,x_1]
haftmann@24155
   161
haftmann@24155
   162
   Moreover, to handle functions that are still waiting for some
haftmann@24155
   163
   arguments we have additionally a list of arguments collected to far
haftmann@24155
   164
   and the number of arguments we're still waiting for.
haftmann@24155
   165
*)
haftmann@24155
   166
wenzelm@25204
   167
datatype Univ =
haftmann@26064
   168
    Const of int * Univ list           (*named (uninterpreted) constants*)
haftmann@25924
   169
  | DFree of string * int              (*free (uninterpreted) dictionary parameters*)
haftmann@27499
   170
  | BVar of int * Univ list            (*bound variables, named*)
haftmann@24155
   171
  | Abs of (int * (Univ list -> Univ)) * Univ list
haftmann@27499
   172
                                       (*abstractions as closures*);
haftmann@24155
   173
haftmann@35371
   174
haftmann@24155
   175
(* constructor functions *)
haftmann@24155
   176
haftmann@25944
   177
fun abss n f = Abs ((n, f), []);
haftmann@25924
   178
fun apps (Abs ((n, f), xs)) ys = let val k = n - length ys in
haftmann@27499
   179
      case int_ord (k, 0)
haftmann@27499
   180
       of EQUAL => f (ys @ xs)
haftmann@27499
   181
        | LESS => let val (zs, ws) = chop (~ k) ys in apps (f (ws @ xs)) zs end
haftmann@27499
   182
        | GREATER => Abs ((k, f), ys @ xs) (*note: reverse convention also for apps!*)
haftmann@27499
   183
      end
haftmann@25924
   184
  | apps (Const (name, xs)) ys = Const (name, ys @ xs)
haftmann@27499
   185
  | apps (BVar (n, xs)) ys = BVar (n, ys @ xs);
haftmann@24155
   186
haftmann@41285
   187
fun same (Const (k, xs), Const (l, ys)) = k = l andalso eq_list same (xs, ys)
haftmann@41285
   188
  | same (DFree (s, k), DFree (t, l)) = s = t andalso k = l
haftmann@41285
   189
  | same (BVar (k, xs), BVar (l, ys)) = k = l andalso eq_list same (xs, ys)
haftmann@41285
   190
  | same _ = false;
haftmann@40971
   191
haftmann@24155
   192
haftmann@24155
   193
(** assembling and compiling ML code from terms **)
haftmann@24155
   194
haftmann@24155
   195
(* abstract ML syntax *)
haftmann@24155
   196
haftmann@24155
   197
infix 9 `$` `$$`;
haftmann@24155
   198
fun e1 `$` e2 = "(" ^ e1 ^ " " ^ e2 ^ ")";
haftmann@25101
   199
fun e `$$` [] = e
haftmann@25101
   200
  | e `$$` es = "(" ^ e ^ " " ^ space_implode " " es ^ ")";
haftmann@24590
   201
fun ml_abs v e = "(fn " ^ v ^ " => " ^ e ^ ")";
haftmann@24155
   202
haftmann@24155
   203
fun ml_cases t cs =
haftmann@24155
   204
  "(case " ^ t ^ " of " ^ space_implode " | " (map (fn (p, t) => p ^ " => " ^ t) cs) ^ ")";
haftmann@25944
   205
fun ml_Let d e = "let\n" ^ d ^ " in " ^ e ^ " end";
haftmann@28337
   206
fun ml_as v t = "(" ^ v ^ " as " ^ t ^ ")";
haftmann@24155
   207
haftmann@28350
   208
fun ml_and [] = "true"
haftmann@28350
   209
  | ml_and [x] = x
haftmann@28350
   210
  | ml_and xs = "(" ^ space_implode " andalso " xs ^ ")";
haftmann@28350
   211
fun ml_if b x y = "(if " ^ b ^ " then " ^ x ^ " else " ^ y ^ ")";
haftmann@28350
   212
haftmann@24155
   213
fun ml_list es = "[" ^ commas es ^ "]";
haftmann@24155
   214
haftmann@24155
   215
fun ml_fundefs ([(name, [([], e)])]) =
haftmann@24155
   216
      "val " ^ name ^ " = " ^ e ^ "\n"
haftmann@24155
   217
  | ml_fundefs (eqs :: eqss) =
haftmann@24155
   218
      let
haftmann@24155
   219
        fun fundef (name, eqs) =
haftmann@24155
   220
          let
haftmann@24155
   221
            fun eqn (es, e) = name ^ " " ^ space_implode " " es ^ " = " ^ e
haftmann@24155
   222
          in space_implode "\n  | " (map eqn eqs) end;
haftmann@24155
   223
      in
haftmann@24155
   224
        (prefix "fun " o fundef) eqs :: map (prefix "and " o fundef) eqss
wenzelm@26931
   225
        |> cat_lines
haftmann@24155
   226
        |> suffix "\n"
haftmann@24155
   227
      end;
haftmann@24155
   228
haftmann@35371
   229
haftmann@25944
   230
(* nbe specific syntax and sandbox communication *)
haftmann@25944
   231
wenzelm@41720
   232
structure Univs = Proof_Data
wenzelm@41720
   233
(
haftmann@39634
   234
  type T = unit -> Univ list -> Univ list
wenzelm@41720
   235
  (* FIXME avoid user error with non-user text *)
haftmann@39643
   236
  fun init _ () = error "Univs"
haftmann@39634
   237
);
haftmann@39634
   238
val put_result = Univs.put;
haftmann@24155
   239
haftmann@24155
   240
local
haftmann@41316
   241
  val prefix =     "Nbe.";
haftmann@41316
   242
  val name_put =   prefix ^ "put_result";
haftmann@41316
   243
  val name_ref =   prefix ^ "univs_ref";
haftmann@41316
   244
  val name_const = prefix ^ "Const";
haftmann@41316
   245
  val name_abss =  prefix ^ "abss";
haftmann@41316
   246
  val name_apps =  prefix ^ "apps";
haftmann@41316
   247
  val name_same =  prefix ^ "same";
haftmann@24155
   248
in
haftmann@24155
   249
haftmann@39634
   250
val univs_cookie = (Univs.get, put_result, name_put);
haftmann@25944
   251
haftmann@28337
   252
fun nbe_fun 0 "" = "nbe_value"
haftmann@28337
   253
  | nbe_fun i c = "c_" ^ translate_string (fn "." => "_" | c => c) c ^ "_" ^ string_of_int i;
haftmann@25101
   254
fun nbe_dict v n = "d_" ^ v ^ "_" ^ string_of_int n;
haftmann@24155
   255
fun nbe_bound v = "v_" ^ v;
haftmann@31888
   256
fun nbe_bound_optional NONE = "_"
haftmann@35486
   257
  | nbe_bound_optional (SOME v) = nbe_bound v;
haftmann@28337
   258
fun nbe_default v = "w_" ^ v;
haftmann@24155
   259
haftmann@25924
   260
(*note: these three are the "turning spots" where proper argument order is established!*)
haftmann@25924
   261
fun nbe_apps t [] = t
haftmann@25924
   262
  | nbe_apps t ts = name_apps `$$` [t, ml_list (rev ts)];
haftmann@28337
   263
fun nbe_apps_local i c ts = nbe_fun i c `$` ml_list (rev ts);
haftmann@28337
   264
fun nbe_apps_constr idx_of c ts =
haftmann@28337
   265
  let
haftmann@28337
   266
    val c' = if !trace then string_of_int (idx_of c) ^ " (*" ^ c ^ "*)"
haftmann@28337
   267
      else string_of_int (idx_of c);
haftmann@28337
   268
  in name_const `$` ("(" ^ c' ^ ", " ^ ml_list (rev ts) ^ ")") end;
haftmann@25924
   269
haftmann@24155
   270
fun nbe_abss 0 f = f `$` ml_list []
haftmann@25944
   271
  | nbe_abss n f = name_abss `$$` [string_of_int n, f];
haftmann@24155
   272
haftmann@41285
   273
fun nbe_same (v1, v2) = "(" ^ name_same ^ " (" ^ nbe_bound v1 ^ ", " ^ nbe_bound v2 ^ "))";
haftmann@28350
   274
haftmann@24155
   275
end;
haftmann@24155
   276
haftmann@28054
   277
open Basic_Code_Thingol;
haftmann@24155
   278
haftmann@35371
   279
haftmann@25865
   280
(* code generation *)
haftmann@24155
   281
haftmann@26064
   282
fun assemble_eqnss idx_of deps eqnss =
haftmann@25944
   283
  let
haftmann@25944
   284
    fun prep_eqns (c, (vs, eqns)) =
haftmann@25944
   285
      let
haftmann@25944
   286
        val dicts = maps (fn (v, sort) => map_index (nbe_dict v o fst) sort) vs;
haftmann@28337
   287
        val num_args = length dicts + ((length o fst o hd) eqns);
haftmann@25944
   288
      in (c, (num_args, (dicts, eqns))) end;
haftmann@25944
   289
    val eqnss' = map prep_eqns eqnss;
haftmann@25101
   290
haftmann@25944
   291
    fun assemble_constapp c dss ts = 
haftmann@25924
   292
      let
haftmann@41366
   293
        val ts' = (maps o map) assemble_dict dss @ ts;
haftmann@25944
   294
      in case AList.lookup (op =) eqnss' c
haftmann@28337
   295
       of SOME (num_args, _) => if num_args <= length ts'
haftmann@28337
   296
            then let val (ts1, ts2) = chop num_args ts'
haftmann@28337
   297
            in nbe_apps (nbe_apps_local 0 c ts1) ts2
haftmann@28337
   298
            end else nbe_apps (nbe_abss num_args (nbe_fun 0 c)) ts'
haftmann@25944
   299
        | NONE => if member (op =) deps c
haftmann@28337
   300
            then nbe_apps (nbe_fun 0 c) ts'
haftmann@28337
   301
            else nbe_apps_constr idx_of c ts'
haftmann@25924
   302
      end
haftmann@41348
   303
    and assemble_classrels classrels =
haftmann@41348
   304
      fold_rev (fn classrel => assemble_constapp classrel [] o single) classrels
haftmann@41366
   305
    and assemble_dict (Dict (classrels, x)) =
haftmann@41366
   306
          assemble_classrels classrels (assemble_plain_dict x)
haftmann@41366
   307
    and assemble_plain_dict (Dict_Const (inst, dss)) =
haftmann@41366
   308
          assemble_constapp inst dss []
haftmann@41366
   309
      | assemble_plain_dict (Dict_Var (v, (n, _))) =
haftmann@41366
   310
          nbe_dict v n
haftmann@25924
   311
haftmann@28350
   312
    fun assemble_iterm constapp =
haftmann@24155
   313
      let
haftmann@28350
   314
        fun of_iterm match_cont t =
haftmann@25944
   315
          let
haftmann@28054
   316
            val (t', ts) = Code_Thingol.unfold_app t
haftmann@28350
   317
          in of_iapp match_cont t' (fold_rev (cons o of_iterm NONE) ts []) end
haftmann@31049
   318
        and of_iapp match_cont (IConst (c, ((_, dss), _))) ts = constapp c dss ts
haftmann@31889
   319
          | of_iapp match_cont (IVar v) ts = nbe_apps (nbe_bound_optional v) ts
haftmann@31724
   320
          | of_iapp match_cont ((v, _) `|=> t) ts =
haftmann@31888
   321
              nbe_apps (nbe_abss 1 (ml_abs (ml_list [nbe_bound_optional v]) (of_iterm NONE t))) ts
haftmann@28350
   322
          | of_iapp match_cont (ICase (((t, _), cs), t0)) ts =
haftmann@28350
   323
              nbe_apps (ml_cases (of_iterm NONE t)
haftmann@28350
   324
                (map (fn (p, t) => (of_iterm NONE p, of_iterm match_cont t)) cs
haftmann@28350
   325
                  @ [("_", case match_cont of SOME s => s | NONE => of_iterm NONE t0)])) ts
haftmann@25944
   326
      in of_iterm end;
haftmann@24155
   327
haftmann@28350
   328
    fun subst_nonlin_vars args =
haftmann@28350
   329
      let
haftmann@28350
   330
        val vs = (fold o Code_Thingol.fold_varnames)
wenzelm@33008
   331
          (fn v => AList.map_default (op =) (v, 0) (Integer.add 1)) args [];
haftmann@28350
   332
        val names = Name.make_context (map fst vs);
haftmann@28350
   333
        fun declare v k ctxt = let val vs = Name.invents ctxt v k
haftmann@28350
   334
          in (vs, fold Name.declare vs ctxt) end;
haftmann@28350
   335
        val (vs_renames, _) = fold_map (fn (v, k) => if k > 1
haftmann@28350
   336
          then declare v (k - 1) #>> (fn vs => (v, vs))
haftmann@28350
   337
          else pair (v, [])) vs names;
haftmann@28350
   338
        val samepairs = maps (fn (v, vs) => map (pair v) vs) vs_renames;
haftmann@28350
   339
        fun subst_vars (t as IConst _) samepairs = (t, samepairs)
haftmann@31889
   340
          | subst_vars (t as IVar NONE) samepairs = (t, samepairs)
haftmann@31889
   341
          | subst_vars (t as IVar (SOME v)) samepairs = (case AList.lookup (op =) samepairs v
haftmann@31889
   342
             of SOME v' => (IVar (SOME v'), AList.delete (op =) v samepairs)
haftmann@28350
   343
              | NONE => (t, samepairs))
haftmann@28350
   344
          | subst_vars (t1 `$ t2) samepairs = samepairs
haftmann@28350
   345
              |> subst_vars t1
haftmann@28350
   346
              ||>> subst_vars t2
haftmann@28350
   347
              |>> (op `$)
haftmann@28350
   348
          | subst_vars (ICase (_, t)) samepairs = subst_vars t samepairs;
haftmann@28350
   349
        val (args', _) = fold_map subst_vars args samepairs;
haftmann@28350
   350
      in (samepairs, args') end;
haftmann@28350
   351
haftmann@28337
   352
    fun assemble_eqn c dicts default_args (i, (args, rhs)) =
haftmann@28337
   353
      let
haftmann@28337
   354
        val is_eval = (c = "");
haftmann@28337
   355
        val default_rhs = nbe_apps_local (i+1) c (dicts @ default_args);
haftmann@28337
   356
        val match_cont = if is_eval then NONE else SOME default_rhs;
haftmann@28350
   357
        val assemble_arg = assemble_iterm
haftmann@28350
   358
          (fn c => fn _ => fn ts => nbe_apps_constr idx_of c ts) NONE;
haftmann@35371
   359
        val assemble_rhs = assemble_iterm assemble_constapp match_cont;
haftmann@28350
   360
        val (samepairs, args') = subst_nonlin_vars args;
haftmann@28350
   361
        val s_args = map assemble_arg args';
haftmann@28350
   362
        val s_rhs = if null samepairs then assemble_rhs rhs
haftmann@41285
   363
          else ml_if (ml_and (map nbe_same samepairs))
haftmann@28350
   364
            (assemble_rhs rhs) default_rhs;
haftmann@28337
   365
        val eqns = if is_eval then
haftmann@28350
   366
            [([ml_list (rev (dicts @ s_args))], s_rhs)]
haftmann@28337
   367
          else
haftmann@28350
   368
            [([ml_list (rev (dicts @ map2 ml_as default_args s_args))], s_rhs),
haftmann@28337
   369
              ([ml_list (rev (dicts @ default_args))], default_rhs)]
haftmann@28337
   370
      in (nbe_fun i c, eqns) end;
haftmann@28337
   371
haftmann@25944
   372
    fun assemble_eqns (c, (num_args, (dicts, eqns))) =
haftmann@25944
   373
      let
haftmann@28337
   374
        val default_args = map nbe_default
haftmann@28337
   375
          (Name.invent_list [] "a" (num_args - length dicts));
haftmann@28337
   376
        val eqns' = map_index (assemble_eqn c dicts default_args) eqns
haftmann@28337
   377
          @ (if c = "" then [] else [(nbe_fun (length eqns) c,
haftmann@28337
   378
            [([ml_list (rev (dicts @ default_args))],
haftmann@28337
   379
              nbe_apps_constr idx_of c (dicts @ default_args))])]);
haftmann@28337
   380
      in (eqns', nbe_abss num_args (nbe_fun 0 c)) end;
haftmann@24155
   381
haftmann@25944
   382
    val (fun_vars, fun_vals) = map_split assemble_eqns eqnss';
haftmann@28337
   383
    val deps_vars = ml_list (map (nbe_fun 0) deps);
haftmann@28337
   384
  in ml_abs deps_vars (ml_Let (ml_fundefs (flat fun_vars)) (ml_list fun_vals)) end;
haftmann@25944
   385
haftmann@35371
   386
haftmann@39636
   387
(* compile equations *)
haftmann@25944
   388
haftmann@39636
   389
fun compile_eqnss thy nbe_program raw_deps [] = []
haftmann@39636
   390
  | compile_eqnss thy nbe_program raw_deps eqnss =
haftmann@24155
   391
      let
wenzelm@43232
   392
        val ctxt = Proof_Context.init_global thy;
haftmann@26064
   393
        val (deps, deps_vals) = split_list (map_filter
haftmann@39636
   394
          (fn dep => Option.map (fn univ => (dep, univ)) (fst ((Graph.get_node nbe_program dep)))) raw_deps);
haftmann@26064
   395
        val idx_of = raw_deps
haftmann@39636
   396
          |> map (fn dep => (dep, snd (Graph.get_node nbe_program dep)))
haftmann@26064
   397
          |> AList.lookup (op =)
haftmann@26064
   398
          |> (fn f => the o f);
haftmann@26064
   399
        val s = assemble_eqnss idx_of deps eqnss;
haftmann@24155
   400
        val cs = map fst eqnss;
haftmann@25944
   401
      in
haftmann@25944
   402
        s
haftmann@32544
   403
        |> traced (fn s => "\n--- code to be evaluated:\n" ^ s)
haftmann@39157
   404
        |> pair ""
haftmann@40092
   405
        |> Code_Runtime.value ctxt univs_cookie
haftmann@25944
   406
        |> (fn f => f deps_vals)
haftmann@25944
   407
        |> (fn univs => cs ~~ univs)
haftmann@25944
   408
      end;
haftmann@25190
   409
wenzelm@30687
   410
haftmann@39636
   411
(* extract equations from statements *)
haftmann@24155
   412
haftmann@37412
   413
fun eqns_of_stmt (_, Code_Thingol.Fun (_, ((_, []), _))) =
haftmann@25101
   414
      []
haftmann@37412
   415
  | eqns_of_stmt (const, Code_Thingol.Fun (_, (((vs, _), eqns), _))) =
haftmann@25101
   416
      [(const, (vs, map fst eqns))]
haftmann@28054
   417
  | eqns_of_stmt (_, Code_Thingol.Datatypecons _) =
haftmann@25101
   418
      []
haftmann@28054
   419
  | eqns_of_stmt (_, Code_Thingol.Datatype _) =
haftmann@25101
   420
      []
haftmann@37422
   421
  | eqns_of_stmt (class, Code_Thingol.Class (_, (v, (super_classes, classparams)))) =
haftmann@25101
   422
      let
haftmann@37359
   423
        val names = map snd super_classes @ map fst classparams;
haftmann@25101
   424
        val params = Name.invent_list [] "d" (length names);
haftmann@25101
   425
        fun mk (k, name) =
haftmann@25101
   426
          (name, ([(v, [])],
haftmann@31889
   427
            [([IConst (class, (([], []), [])) `$$ map (IVar o SOME) params],
haftmann@31889
   428
              IVar (SOME (nth params k)))]));
haftmann@25101
   429
      in map_index mk names end
haftmann@28054
   430
  | eqns_of_stmt (_, Code_Thingol.Classrel _) =
haftmann@25101
   431
      []
haftmann@28054
   432
  | eqns_of_stmt (_, Code_Thingol.Classparam _) =
haftmann@25101
   433
      []
haftmann@37424
   434
  | eqns_of_stmt (inst, Code_Thingol.Classinst ((class, (_, arity_args)), (super_instances, (classparam_instances, _)))) =
haftmann@37359
   435
      [(inst, (arity_args, [([], IConst (class, (([], []), [])) `$$
haftmann@37359
   436
        map (fn (_, (_, (inst, dss))) => IConst (inst, (([], dss), []))) super_instances
haftmann@37359
   437
        @ map (IConst o snd o fst) classparam_instances)]))];
haftmann@25101
   438
haftmann@39636
   439
haftmann@39636
   440
(* compile whole programs *)
haftmann@39636
   441
haftmann@39643
   442
fun ensure_const_idx name (nbe_program, (maxidx, idx_tab)) =
haftmann@39643
   443
  if can (Graph.get_node nbe_program) name
haftmann@39643
   444
  then (nbe_program, (maxidx, idx_tab))
haftmann@39643
   445
  else (Graph.new_node (name, (NONE, maxidx)) nbe_program,
haftmann@39643
   446
    (maxidx + 1, Inttab.update_new (maxidx, name) idx_tab));
haftmann@39643
   447
haftmann@39636
   448
fun compile_stmts thy stmts_deps =
haftmann@24155
   449
  let
haftmann@25101
   450
    val names = map (fst o fst) stmts_deps;
haftmann@25101
   451
    val names_deps = map (fn ((name, _), deps) => (name, deps)) stmts_deps;
haftmann@25101
   452
    val eqnss = maps (eqns_of_stmt o fst) stmts_deps;
haftmann@26064
   453
    val refl_deps = names_deps
haftmann@25190
   454
      |> maps snd
haftmann@25190
   455
      |> distinct (op =)
haftmann@26064
   456
      |> fold (insert (op =)) names;
haftmann@39636
   457
    fun compile nbe_program = eqnss
haftmann@39636
   458
      |> compile_eqnss thy nbe_program refl_deps
haftmann@39636
   459
      |> rpair nbe_program;
haftmann@25101
   460
  in
haftmann@39643
   461
    fold ensure_const_idx refl_deps
haftmann@26064
   462
    #> apfst (fold (fn (name, deps) => fold (curry Graph.add_edge name) deps) names_deps
haftmann@26064
   463
      #> compile
haftmann@26064
   464
      #-> fold (fn (name, univ) => (Graph.map_node name o apfst) (K (SOME univ))))
haftmann@25101
   465
  end;
haftmann@25101
   466
haftmann@39636
   467
fun compile_program thy program =
haftmann@25101
   468
  let
haftmann@39636
   469
    fun add_stmts names (nbe_program, (maxidx, idx_tab)) = if exists ((can o Graph.get_node) nbe_program) names
haftmann@39636
   470
      then (nbe_program, (maxidx, idx_tab))
haftmann@39636
   471
      else (nbe_program, (maxidx, idx_tab))
haftmann@39636
   472
        |> compile_stmts thy (map (fn name => ((name, Graph.get_node program name),
haftmann@27103
   473
          Graph.imm_succs program name)) names);
haftmann@28706
   474
  in
haftmann@28706
   475
    fold_rev add_stmts (Graph.strong_conn program)
haftmann@28706
   476
  end;
haftmann@25101
   477
haftmann@25944
   478
haftmann@25944
   479
(** evaluation **)
haftmann@25944
   480
haftmann@39636
   481
(* term evaluation by compilation *)
haftmann@25944
   482
haftmann@39636
   483
fun compile_term thy nbe_program deps (vs : (string * sort) list, t) =
haftmann@25924
   484
  let 
haftmann@25924
   485
    val dict_frees = maps (fn (v, sort) => map_index (curry DFree v o fst) sort) vs;
haftmann@25924
   486
  in
haftmann@31064
   487
    ("", (vs, [([], t)]))
haftmann@39636
   488
    |> singleton (compile_eqnss thy nbe_program deps)
haftmann@25924
   489
    |> snd
haftmann@31064
   490
    |> (fn t => apps t (rev dict_frees))
haftmann@25924
   491
  end;
haftmann@24155
   492
haftmann@35371
   493
haftmann@39636
   494
(* reconstruction *)
haftmann@24155
   495
haftmann@30947
   496
fun typ_of_itype program vs (ityco `%% itys) =
haftmann@30947
   497
      let
haftmann@30947
   498
        val Code_Thingol.Datatype (tyco, _) = Graph.get_node program ityco;
haftmann@30947
   499
      in Type (tyco, map (typ_of_itype program vs) itys) end
haftmann@30947
   500
  | typ_of_itype program vs (ITyVar v) =
haftmann@30947
   501
      let
haftmann@30947
   502
        val sort = (the o AList.lookup (op =) vs) v;
haftmann@30947
   503
      in TFree ("'" ^ v, sort) end;
haftmann@30947
   504
haftmann@28663
   505
fun term_of_univ thy program idx_tab t =
haftmann@24155
   506
  let
haftmann@25101
   507
    fun take_until f [] = []
haftmann@25101
   508
      | take_until f (x::xs) = if f x then [] else x :: take_until f xs;
haftmann@28663
   509
    fun is_dict (Const (idx, _)) = (case (Graph.get_node program o the o Inttab.lookup idx_tab) idx
haftmann@28663
   510
         of Code_Thingol.Class _ => true
haftmann@28663
   511
          | Code_Thingol.Classrel _ => true
haftmann@28663
   512
          | Code_Thingol.Classinst _ => true
haftmann@28663
   513
          | _ => false)
haftmann@25101
   514
      | is_dict (DFree _) = true
haftmann@25101
   515
      | is_dict _ = false;
haftmann@28663
   516
    fun const_of_idx idx = (case (Graph.get_node program o the o Inttab.lookup idx_tab) idx
haftmann@35371
   517
     of Code_Thingol.Fun (c, _) => c
haftmann@35371
   518
      | Code_Thingol.Datatypecons (c, _) => c
haftmann@35371
   519
      | Code_Thingol.Classparam (c, _) => c);
haftmann@24155
   520
    fun of_apps bounds (t, ts) =
haftmann@24155
   521
      fold_map (of_univ bounds) ts
haftmann@24155
   522
      #>> (fn ts' => list_comb (t, rev ts'))
haftmann@26064
   523
    and of_univ bounds (Const (idx, ts)) typidx =
haftmann@24155
   524
          let
haftmann@25101
   525
            val ts' = take_until is_dict ts;
haftmann@28663
   526
            val c = const_of_idx idx;
haftmann@31957
   527
            val T = map_type_tvar (fn ((v, i), _) =>
wenzelm@37153
   528
              Type_Infer.param typidx (v ^ string_of_int i, []))
haftmann@31957
   529
                (Sign.the_const_type thy c);
haftmann@29959
   530
            val typidx' = typidx + 1;
haftmann@31957
   531
          in of_apps bounds (Term.Const (c, T), ts') typidx' end
haftmann@27499
   532
      | of_univ bounds (BVar (n, ts)) typidx =
haftmann@27499
   533
          of_apps bounds (Bound (bounds - n - 1), ts) typidx
haftmann@24155
   534
      | of_univ bounds (t as Abs _) typidx =
haftmann@24155
   535
          typidx
haftmann@25924
   536
          |> of_univ (bounds + 1) (apps t [BVar (bounds, [])])
haftmann@24155
   537
          |-> (fn t' => pair (Term.Abs ("u", dummyT, t')))
haftmann@24155
   538
  in of_univ 0 t 0 |> fst end;
haftmann@24155
   539
haftmann@35371
   540
haftmann@39636
   541
(* evaluation with type reconstruction *)
haftmann@39636
   542
haftmann@39636
   543
fun eval_term thy program (nbe_program, idx_tab) ((vs0, (vs, ty)), t) deps =
haftmann@39636
   544
  let
haftmann@39636
   545
    val ctxt = Syntax.init_pretty_global thy;
haftmann@39636
   546
    val string_of_term = Syntax.string_of_term (Config.put show_types true ctxt);
haftmann@39636
   547
    val ty' = typ_of_itype program vs0 ty;
haftmann@39636
   548
    fun type_infer t = singleton
haftmann@39636
   549
      (Type_Infer.infer_types ctxt (try (Type.strip_sorts o Sign.the_const_type thy)) (K NONE))
haftmann@39636
   550
      (Type.constraint ty' t);
haftmann@39636
   551
    fun check_tvars t =
haftmann@39636
   552
      if null (Term.add_tvars t []) then t
haftmann@39636
   553
      else error ("Illegal schematic type variables in normalized term: " ^ string_of_term t);
haftmann@39636
   554
  in
haftmann@39636
   555
    compile_term thy nbe_program deps (vs, t)
haftmann@39636
   556
    |> term_of_univ thy program idx_tab
haftmann@39636
   557
    |> traced (fn t => "Normalized:\n" ^ string_of_term t)
haftmann@39636
   558
    |> type_infer
haftmann@39636
   559
    |> traced (fn t => "Types inferred:\n" ^ string_of_term t)
haftmann@39636
   560
    |> check_tvars
haftmann@39636
   561
    |> traced (fn _ => "---\n")
haftmann@39636
   562
  end;
haftmann@39636
   563
haftmann@39682
   564
haftmann@25101
   565
(* function store *)
haftmann@25101
   566
haftmann@34160
   567
structure Nbe_Functions = Code_Data
haftmann@25101
   568
(
haftmann@35486
   569
  type T = (Univ option * int) Graph.T * (int * string Inttab.table);
haftmann@35486
   570
  val empty = (Graph.empty, (0, Inttab.empty));
haftmann@25101
   571
);
haftmann@25101
   572
haftmann@39643
   573
fun compile ignore_cache thy program =
haftmann@26064
   574
  let
haftmann@39636
   575
    val (nbe_program, (_, idx_tab)) =
haftmann@39643
   576
      Nbe_Functions.change (if ignore_cache then NONE else SOME thy)
haftmann@39643
   577
        (compile_program thy program);
haftmann@39636
   578
  in (nbe_program, idx_tab) end;
haftmann@24155
   579
haftmann@32544
   580
haftmann@39640
   581
(* dynamic evaluation oracle *)
haftmann@24155
   582
haftmann@30947
   583
fun mk_equals thy lhs raw_rhs =
haftmann@26747
   584
  let
haftmann@30947
   585
    val ty = Thm.typ_of (Thm.ctyp_of_term lhs);
haftmann@30947
   586
    val eq = Thm.cterm_of thy (Term.Const ("==", ty --> ty --> propT));
haftmann@30947
   587
    val rhs = Thm.cterm_of thy raw_rhs;
haftmann@30947
   588
  in Thm.mk_binop eq lhs rhs end;
haftmann@26747
   589
haftmann@38908
   590
val (_, raw_oracle) = Context.>>> (Context.map_theory_result
haftmann@39643
   591
  (Thm.add_oracle (Binding.name "normalization_by_evaluation",
haftmann@39643
   592
    fn (thy, program, nbe_program_idx_tab, vsp_ty_t, deps, ct) =>
haftmann@39643
   593
      mk_equals thy ct (eval_term thy program nbe_program_idx_tab vsp_ty_t deps))));
haftmann@30942
   594
haftmann@39643
   595
fun oracle thy program nbe_program_idx_tab vsp_ty_t deps ct =
haftmann@39643
   596
  raw_oracle (thy, program, nbe_program_idx_tab, vsp_ty_t, deps, ct);
haftmann@30942
   597
haftmann@41488
   598
fun dynamic_conv thy = lift_triv_classes_conv thy (Code_Thingol.dynamic_conv thy
haftmann@41488
   599
    (K (fn program => oracle thy program (compile false thy program))));
haftmann@24155
   600
haftmann@41432
   601
fun dynamic_value thy = lift_triv_classes_rew thy
haftmann@41432
   602
  (Code_Thingol.dynamic_value thy I
haftmann@39830
   603
    (K (fn program => eval_term thy program (compile false thy program))));
haftmann@39643
   604
haftmann@41432
   605
fun static_conv thy consts =
haftmann@41432
   606
  lift_triv_classes_conv thy (Code_Thingol.static_conv thy consts
haftmann@41594
   607
    (K (fn program => fn _ => let val nbe_program = compile true thy program
haftmann@41594
   608
      in oracle thy program nbe_program end)));
haftmann@24839
   609
haftmann@41432
   610
fun static_value thy consts = lift_triv_classes_rew thy
haftmann@41432
   611
  (Code_Thingol.static_value thy I consts
haftmann@41594
   612
    (K (fn program => fn _ => let val nbe_program = compile true thy program
haftmann@41594
   613
      in eval_term thy program (compile false thy program) end)));
haftmann@41432
   614
haftmann@35371
   615
haftmann@39640
   616
(** setup **)
haftmann@24155
   617
wenzelm@43232
   618
val setup = Value.add_evaluator ("nbe", dynamic_value o Proof_Context.theory_of);
haftmann@24155
   619
haftmann@24155
   620
end;
haftmann@28337
   621