src/Tools/nbe.ML
author haftmann
Sat, 25 Jan 2014 23:50:49 +0100
changeset 55147 bce3dbc11f95
parent 55043 acefda71629b
child 55150 0940309ed8f1
permissions -rw-r--r--
prefer explicit code symbol type over ad-hoc name mangling
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
24590
733120d04233 delayed evaluation
haftmann
parents: 24508
diff changeset
     1
(*  Title:      Tools/nbe.ML
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
     2
    Authors:    Klaus Aehlig, LMU Muenchen; Tobias Nipkow, Florian Haftmann, TU Muenchen
d86867645f4f nbe improved
haftmann
parents:
diff changeset
     3
24839
199c48ec5a09 step towards proper purge operation
haftmann
parents: 24713
diff changeset
     4
Normalization by evaluation, based on generic code generator.
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
     5
*)
d86867645f4f nbe improved
haftmann
parents:
diff changeset
     6
d86867645f4f nbe improved
haftmann
parents:
diff changeset
     7
signature NBE =
d86867645f4f nbe improved
haftmann
parents:
diff changeset
     8
sig
41247
c5cb19ecbd41 avoid slightly odd Conv.tap_thy
haftmann
parents: 41184
diff changeset
     9
  val dynamic_conv: theory -> conv
41184
5c6f44d22f51 simplified evaluation function names
haftmann
parents: 41118
diff changeset
    10
  val dynamic_value: theory -> term -> term
5c6f44d22f51 simplified evaluation function names
haftmann
parents: 41118
diff changeset
    11
  val static_conv: theory -> string list -> conv
5c6f44d22f51 simplified evaluation function names
haftmann
parents: 41118
diff changeset
    12
  val static_value: theory -> string list -> term -> term
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
    13
25204
36cf92f63a44 replaced Secure.evaluate by ML_Context.evaluate;
wenzelm
parents: 25190
diff changeset
    14
  datatype Univ =
26064
65585de05a66 using integers for pattern matching
haftmann
parents: 26011
diff changeset
    15
      Const of int * Univ list               (*named (uninterpreted) constants*)
25924
f974a1c64348 improved implementation
haftmann
parents: 25865
diff changeset
    16
    | DFree of string * int                  (*free (uninterpreted) dictionary parameters*)
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
    17
    | BVar of int * Univ list
28350
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
    18
    | Abs of (int * (Univ list -> Univ)) * Univ list
25924
f974a1c64348 improved implementation
haftmann
parents: 25865
diff changeset
    19
  val apps: Univ -> Univ list -> Univ        (*explicit applications*)
25944
haftmann
parents: 25935
diff changeset
    20
  val abss: int -> (Univ list -> Univ) -> Univ
28337
93964076e7b8 case default fallback for NBE
haftmann
parents: 28296
diff changeset
    21
                                             (*abstractions as closures*)
41037
6d6f23b3a879 removed experimental equality checking of closures; acknowledge underapproximation of equality in function name
haftmann
parents: 40730
diff changeset
    22
  val same: Univ * Univ -> bool
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
    23
39388
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39290
diff changeset
    24
  val put_result: (unit -> Univ list -> Univ list) -> Proof.context -> Proof.context
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32544
diff changeset
    25
  val trace: bool Unsynchronized.ref
25924
f974a1c64348 improved implementation
haftmann
parents: 25865
diff changeset
    26
26747
f32fa5f5bdd1 moved 'trivial classes' to foundation of code generator
haftmann
parents: 26739
diff changeset
    27
  val setup: theory -> theory
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    28
  val add_const_alias: thm -> theory -> theory
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
    29
end;
d86867645f4f nbe improved
haftmann
parents:
diff changeset
    30
d86867645f4f nbe improved
haftmann
parents:
diff changeset
    31
structure Nbe: NBE =
d86867645f4f nbe improved
haftmann
parents:
diff changeset
    32
struct
d86867645f4f nbe improved
haftmann
parents:
diff changeset
    33
d86867645f4f nbe improved
haftmann
parents:
diff changeset
    34
(* generic non-sense *)
d86867645f4f nbe improved
haftmann
parents:
diff changeset
    35
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32544
diff changeset
    36
val trace = Unsynchronized.ref false;
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    37
fun traced f x = if !trace then (tracing (f x); x) else x;
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
    38
d86867645f4f nbe improved
haftmann
parents:
diff changeset
    39
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    40
(** certificates and oracle for "trivial type classes" **)
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    41
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33002
diff changeset
    42
structure Triv_Class_Data = Theory_Data
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    43
(
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    44
  type T = (class * thm) list;
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    45
  val empty = [];
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    46
  val extend = I;
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33002
diff changeset
    47
  fun merge data : T = AList.merge (op =) (K true) data;
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    48
);
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    49
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    50
fun add_const_alias thm thy =
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    51
  let
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    52
    val (ofclass, eqn) = case try Logic.dest_equals (Thm.prop_of thm)
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    53
     of SOME ofclass_eq => ofclass_eq
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    54
      | _ => error ("Bad certificate: " ^ Display.string_of_thm_global thy thm);
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    55
    val (T, class) = case try Logic.dest_of_class ofclass
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    56
     of SOME T_class => T_class
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    57
      | _ => error ("Bad certificate: " ^ Display.string_of_thm_global thy thm);
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    58
    val tvar = case try Term.dest_TVar T
51685
385ef6706252 more standard module name Axclass (according to file name);
wenzelm
parents: 48072
diff changeset
    59
     of SOME (tvar as (_, sort)) => if null (filter (can (Axclass.get_info thy)) sort)
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    60
          then tvar
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    61
          else error ("Bad sort: " ^ Display.string_of_thm_global thy thm)
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    62
      | _ => error ("Bad type: " ^ Display.string_of_thm_global thy thm);
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    63
    val _ = if Term.add_tvars eqn [] = [tvar] then ()
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    64
      else error ("Inconsistent type: " ^ Display.string_of_thm_global thy thm);
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    65
    val lhs_rhs = case try Logic.dest_equals eqn
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    66
     of SOME lhs_rhs => lhs_rhs
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    67
      | _ => error ("Not an equation: " ^ Syntax.string_of_term_global thy eqn);
51685
385ef6706252 more standard module name Axclass (according to file name);
wenzelm
parents: 48072
diff changeset
    68
    val c_c' = case try (pairself (Axclass.unoverload_const thy o dest_Const)) lhs_rhs
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    69
     of SOME c_c' => c_c'
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    70
      | _ => error ("Not an equation with two constants: "
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    71
          ^ Syntax.string_of_term_global thy eqn);
51685
385ef6706252 more standard module name Axclass (according to file name);
wenzelm
parents: 48072
diff changeset
    72
    val _ = if the_list (Axclass.class_of_param thy (snd c_c')) = [class] then ()
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    73
      else error ("Inconsistent class: " ^ Display.string_of_thm_global thy thm);
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    74
  in Triv_Class_Data.map (AList.update (op =) (class, thm)) thy end;
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    75
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    76
local
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    77
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    78
val get_triv_classes = map fst o Triv_Class_Data.get;
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    79
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    80
val (_, triv_of_class) = Context.>>> (Context.map_theory_result
43619
3803869014aa proper @{binding} antiquotations (relevant for formal references);
wenzelm
parents: 43329
diff changeset
    81
  (Thm.add_oracle (@{binding triv_of_class}, fn (thy, T, class) =>
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    82
    Thm.cterm_of thy (Logic.mk_of_class (T, class)))));
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    83
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    84
in
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    85
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    86
fun lift_triv_classes_conv thy conv ct =
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    87
  let
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    88
    val algebra = Sign.classes_of thy;
36982
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
    89
    val certT = Thm.ctyp_of thy;
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    90
    val triv_classes = get_triv_classes thy;
36982
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
    91
    fun additional_classes sort = filter_out (fn class => Sorts.sort_le algebra (sort, [class])) triv_classes;
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
    92
    fun mk_entry (v, sort) =
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
    93
      let
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
    94
        val T = TFree (v, sort);
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
    95
        val cT = certT T;
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
    96
        val triv_sort = additional_classes sort;
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
    97
      in
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
    98
        (v, (Sorts.inter_sort algebra (sort, triv_sort),
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
    99
          (cT, AList.make (fn class => Thm.of_class (cT, class)) sort
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   100
            @ AList.make (fn class => triv_of_class (thy, T, class)) triv_sort)))
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   101
      end;
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   102
    val vs_tab = map mk_entry (Term.add_tfrees (Thm.term_of ct) []);
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   103
    fun instantiate thm =
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   104
      let
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   105
        val cert_tvars = map (certT o TVar) (Term.add_tvars
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   106
          ((fst o Logic.dest_equals o Logic.strip_imp_concl o Thm.prop_of) thm) []);
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   107
        val instantiation =
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   108
          map2 (fn cert_tvar => fn (_, (_, (cT, _))) => (cert_tvar, cT)) cert_tvars vs_tab;
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   109
      in Thm.instantiate (instantiation, []) thm end;
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   110
    fun of_class (TFree (v, _), class) =
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   111
          the (AList.lookup (op =) ((snd o snd o the o AList.lookup (op =) vs_tab) v) class)
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   112
      | of_class (T, _) = error ("Bad type " ^ Syntax.string_of_typ_global thy T);
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   113
    fun strip_of_class thm =
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   114
      let
36982
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   115
        val prems_of_class = Thm.prop_of thm
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   116
          |> Logic.strip_imp_prems
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   117
          |> map (Logic.dest_of_class #> of_class);
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   118
      in fold Thm.elim_implies prems_of_class thm end;
36770
c3a04614c710 reactivated Thm.legacy_unconstrainT for Nbe.lift_triv_classes_conv;
wenzelm
parents: 36768
diff changeset
   119
  in
c3a04614c710 reactivated Thm.legacy_unconstrainT for Nbe.lift_triv_classes_conv;
wenzelm
parents: 36768
diff changeset
   120
    ct
36982
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   121
    |> (Drule.cterm_fun o map_types o map_type_tfree)
47609
b3dab1892cda dropped dead code
haftmann
parents: 47573
diff changeset
   122
        (fn (v, _) => TFree (v, (fst o the o AList.lookup (op =) vs_tab) v))
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   123
    |> conv
36982
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   124
    |> Thm.strip_shyps
35845
e5980f0ad025 renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents: 35500
diff changeset
   125
    |> Thm.varifyT_global
36982
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   126
    |> Thm.unconstrainT
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   127
    |> instantiate
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   128
    |> strip_of_class
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   129
  end;
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   130
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   131
fun lift_triv_classes_rew thy rew t =
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   132
  let
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   133
    val algebra = Sign.classes_of thy;
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   134
    val triv_classes = get_triv_classes thy;
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   135
    val vs = Term.add_tfrees t [];
52473
a2407f62a29f do not choke on type variables emerging during rewriting
haftmann
parents: 51685
diff changeset
   136
  in
a2407f62a29f do not choke on type variables emerging during rewriting
haftmann
parents: 51685
diff changeset
   137
    t
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   138
    |> (map_types o map_type_tfree)
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   139
        (fn (v, sort) => TFree (v, Sorts.inter_sort algebra (sort, triv_classes)))
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   140
    |> rew
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   141
    |> (map_types o map_type_tfree)
52473
a2407f62a29f do not choke on type variables emerging during rewriting
haftmann
parents: 51685
diff changeset
   142
        (fn (v, sort) => TFree (v, the_default sort (AList.lookup (op =) vs v)))
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   143
  end;
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   144
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   145
end;
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   146
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   147
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   148
(** the semantic universe **)
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   149
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   150
(*
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   151
   Functions are given by their semantical function value. To avoid
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   152
   trouble with the ML-type system, these functions have the most
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   153
   generic type, that is "Univ list -> Univ". The calling convention is
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   154
   that the arguments come as a list, the last argument first. In
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   155
   other words, a function call that usually would look like
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   156
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   157
   f x_1 x_2 ... x_n   or   f(x_1,x_2, ..., x_n)
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   158
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   159
   would be in our convention called as
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   160
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   161
              f [x_n,..,x_2,x_1]
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   162
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   163
   Moreover, to handle functions that are still waiting for some
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   164
   arguments we have additionally a list of arguments collected to far
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   165
   and the number of arguments we're still waiting for.
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   166
*)
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   167
25204
36cf92f63a44 replaced Secure.evaluate by ML_Context.evaluate;
wenzelm
parents: 25190
diff changeset
   168
datatype Univ =
26064
65585de05a66 using integers for pattern matching
haftmann
parents: 26011
diff changeset
   169
    Const of int * Univ list           (*named (uninterpreted) constants*)
25924
f974a1c64348 improved implementation
haftmann
parents: 25865
diff changeset
   170
  | DFree of string * int              (*free (uninterpreted) dictionary parameters*)
27499
150558266831 clarified code
haftmann
parents: 27264
diff changeset
   171
  | BVar of int * Univ list            (*bound variables, named*)
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   172
  | Abs of (int * (Univ list -> Univ)) * Univ list
27499
150558266831 clarified code
haftmann
parents: 27264
diff changeset
   173
                                       (*abstractions as closures*);
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   174
35371
6c92eb394e3c tuned whitespace
haftmann
parents: 34251
diff changeset
   175
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   176
(* constructor functions *)
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   177
25944
haftmann
parents: 25935
diff changeset
   178
fun abss n f = Abs ((n, f), []);
25924
f974a1c64348 improved implementation
haftmann
parents: 25865
diff changeset
   179
fun apps (Abs ((n, f), xs)) ys = let val k = n - length ys in
27499
150558266831 clarified code
haftmann
parents: 27264
diff changeset
   180
      case int_ord (k, 0)
150558266831 clarified code
haftmann
parents: 27264
diff changeset
   181
       of EQUAL => f (ys @ xs)
150558266831 clarified code
haftmann
parents: 27264
diff changeset
   182
        | LESS => let val (zs, ws) = chop (~ k) ys in apps (f (ws @ xs)) zs end
150558266831 clarified code
haftmann
parents: 27264
diff changeset
   183
        | GREATER => Abs ((k, f), ys @ xs) (*note: reverse convention also for apps!*)
150558266831 clarified code
haftmann
parents: 27264
diff changeset
   184
      end
25924
f974a1c64348 improved implementation
haftmann
parents: 25865
diff changeset
   185
  | apps (Const (name, xs)) ys = Const (name, ys @ xs)
27499
150558266831 clarified code
haftmann
parents: 27264
diff changeset
   186
  | apps (BVar (n, xs)) ys = BVar (n, ys @ xs);
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   187
41037
6d6f23b3a879 removed experimental equality checking of closures; acknowledge underapproximation of equality in function name
haftmann
parents: 40730
diff changeset
   188
fun same (Const (k, xs), Const (l, ys)) = k = l andalso eq_list same (xs, ys)
6d6f23b3a879 removed experimental equality checking of closures; acknowledge underapproximation of equality in function name
haftmann
parents: 40730
diff changeset
   189
  | same (DFree (s, k), DFree (t, l)) = s = t andalso k = l
6d6f23b3a879 removed experimental equality checking of closures; acknowledge underapproximation of equality in function name
haftmann
parents: 40730
diff changeset
   190
  | same (BVar (k, xs), BVar (l, ys)) = k = l andalso eq_list same (xs, ys)
6d6f23b3a879 removed experimental equality checking of closures; acknowledge underapproximation of equality in function name
haftmann
parents: 40730
diff changeset
   191
  | same _ = false;
40730
2aa0390a2da7 nbe decides equality of abstractions by extensionality
haftmann
parents: 40564
diff changeset
   192
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   193
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   194
(** assembling and compiling ML code from terms **)
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   195
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   196
(* abstract ML syntax *)
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   197
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   198
infix 9 `$` `$$`;
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   199
fun e1 `$` e2 = "(" ^ e1 ^ " " ^ e2 ^ ")";
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   200
fun e `$$` [] = e
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   201
  | e `$$` es = "(" ^ e ^ " " ^ space_implode " " es ^ ")";
24590
733120d04233 delayed evaluation
haftmann
parents: 24508
diff changeset
   202
fun ml_abs v e = "(fn " ^ v ^ " => " ^ e ^ ")";
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   203
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   204
fun ml_cases t cs =
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   205
  "(case " ^ t ^ " of " ^ space_implode " | " (map (fn (p, t) => p ^ " => " ^ t) cs) ^ ")";
25944
haftmann
parents: 25935
diff changeset
   206
fun ml_Let d e = "let\n" ^ d ^ " in " ^ e ^ " end";
28337
93964076e7b8 case default fallback for NBE
haftmann
parents: 28296
diff changeset
   207
fun ml_as v t = "(" ^ v ^ " as " ^ t ^ ")";
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   208
28350
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   209
fun ml_and [] = "true"
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   210
  | ml_and [x] = x
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   211
  | ml_and xs = "(" ^ space_implode " andalso " xs ^ ")";
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   212
fun ml_if b x y = "(if " ^ b ^ " then " ^ x ^ " else " ^ y ^ ")";
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   213
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   214
fun ml_list es = "[" ^ commas es ^ "]";
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   215
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   216
fun ml_fundefs ([(name, [([], e)])]) =
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   217
      "val " ^ name ^ " = " ^ e ^ "\n"
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   218
  | ml_fundefs (eqs :: eqss) =
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   219
      let
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   220
        fun fundef (name, eqs) =
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   221
          let
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   222
            fun eqn (es, e) = name ^ " " ^ space_implode " " es ^ " = " ^ e
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   223
          in space_implode "\n  | " (map eqn eqs) end;
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   224
      in
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   225
        (prefix "fun " o fundef) eqs :: map (prefix "and " o fundef) eqss
26931
aa226d8405a8 cat_lines;
wenzelm
parents: 26747
diff changeset
   226
        |> cat_lines
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   227
        |> suffix "\n"
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   228
      end;
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   229
35371
6c92eb394e3c tuned whitespace
haftmann
parents: 34251
diff changeset
   230
25944
haftmann
parents: 25935
diff changeset
   231
(* nbe specific syntax and sandbox communication *)
haftmann
parents: 25935
diff changeset
   232
41472
f6ab14e61604 misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents: 41346
diff changeset
   233
structure Univs = Proof_Data
f6ab14e61604 misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents: 41346
diff changeset
   234
(
39388
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39290
diff changeset
   235
  type T = unit -> Univ list -> Univ list
41472
f6ab14e61604 misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents: 41346
diff changeset
   236
  (* FIXME avoid user error with non-user text *)
39399
267235a14938 static nbe conversion
haftmann
parents: 39396
diff changeset
   237
  fun init _ () = error "Univs"
39388
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39290
diff changeset
   238
);
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39290
diff changeset
   239
val put_result = Univs.put;
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   240
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   241
local
41068
7e643e07be7f tuned whitespace
haftmann
parents: 41037
diff changeset
   242
  val prefix =     "Nbe.";
7e643e07be7f tuned whitespace
haftmann
parents: 41037
diff changeset
   243
  val name_put =   prefix ^ "put_result";
7e643e07be7f tuned whitespace
haftmann
parents: 41037
diff changeset
   244
  val name_const = prefix ^ "Const";
7e643e07be7f tuned whitespace
haftmann
parents: 41037
diff changeset
   245
  val name_abss =  prefix ^ "abss";
7e643e07be7f tuned whitespace
haftmann
parents: 41037
diff changeset
   246
  val name_apps =  prefix ^ "apps";
7e643e07be7f tuned whitespace
haftmann
parents: 41037
diff changeset
   247
  val name_same =  prefix ^ "same";
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   248
in
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   249
39388
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39290
diff changeset
   250
val univs_cookie = (Univs.get, put_result, name_put);
25944
haftmann
parents: 25935
diff changeset
   251
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   252
val sloppy_name = Code_Symbol.base_name;
55043
acefda71629b prefer indexes user for pattern matching to print concrete names for symbols, do not rely on printable unique identifiers
haftmann
parents: 54889
diff changeset
   253
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   254
fun nbe_fun idx_of 0 (Code_Symbol.Constant "") = "nbe_value"
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   255
  | nbe_fun idx_of i sym = "c_" ^ string_of_int (idx_of sym)
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   256
      ^ "_" ^ sloppy_name sym ^ "_" ^ string_of_int i;
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   257
fun nbe_dict v n = "d_" ^ v ^ "_" ^ string_of_int n;
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   258
fun nbe_bound v = "v_" ^ v;
31888
626c075fd457 variable names in abstractions are optional
haftmann
parents: 31724
diff changeset
   259
fun nbe_bound_optional NONE = "_"
35500
118cda173627 dropped superfluous naming
haftmann
parents: 35371
diff changeset
   260
  | nbe_bound_optional (SOME v) = nbe_bound v;
28337
93964076e7b8 case default fallback for NBE
haftmann
parents: 28296
diff changeset
   261
fun nbe_default v = "w_" ^ v;
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   262
25924
f974a1c64348 improved implementation
haftmann
parents: 25865
diff changeset
   263
(*note: these three are the "turning spots" where proper argument order is established!*)
f974a1c64348 improved implementation
haftmann
parents: 25865
diff changeset
   264
fun nbe_apps t [] = t
f974a1c64348 improved implementation
haftmann
parents: 25865
diff changeset
   265
  | nbe_apps t ts = name_apps `$$` [t, ml_list (rev ts)];
55043
acefda71629b prefer indexes user for pattern matching to print concrete names for symbols, do not rely on printable unique identifiers
haftmann
parents: 54889
diff changeset
   266
fun nbe_apps_local idx_of i c ts = nbe_fun idx_of i c `$` ml_list (rev ts);
28337
93964076e7b8 case default fallback for NBE
haftmann
parents: 28296
diff changeset
   267
fun nbe_apps_constr idx_of c ts =
93964076e7b8 case default fallback for NBE
haftmann
parents: 28296
diff changeset
   268
  let
55043
acefda71629b prefer indexes user for pattern matching to print concrete names for symbols, do not rely on printable unique identifiers
haftmann
parents: 54889
diff changeset
   269
    val c' = if !trace then string_of_int (idx_of c) ^ " (*" ^ sloppy_name c ^ "*)"
28337
93964076e7b8 case default fallback for NBE
haftmann
parents: 28296
diff changeset
   270
      else string_of_int (idx_of c);
93964076e7b8 case default fallback for NBE
haftmann
parents: 28296
diff changeset
   271
  in name_const `$` ("(" ^ c' ^ ", " ^ ml_list (rev ts) ^ ")") end;
25924
f974a1c64348 improved implementation
haftmann
parents: 25865
diff changeset
   272
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   273
fun nbe_abss 0 f = f `$` ml_list []
25944
haftmann
parents: 25935
diff changeset
   274
  | nbe_abss n f = name_abss `$$` [string_of_int n, f];
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   275
41037
6d6f23b3a879 removed experimental equality checking of closures; acknowledge underapproximation of equality in function name
haftmann
parents: 40730
diff changeset
   276
fun nbe_same (v1, v2) = "(" ^ name_same ^ " (" ^ nbe_bound v1 ^ ", " ^ nbe_bound v2 ^ "))";
28350
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   277
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   278
end;
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   279
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents: 27609
diff changeset
   280
open Basic_Code_Thingol;
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   281
35371
6c92eb394e3c tuned whitespace
haftmann
parents: 34251
diff changeset
   282
25865
a141d6bfd398 tuned comment
haftmann
parents: 25204
diff changeset
   283
(* code generation *)
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   284
26064
65585de05a66 using integers for pattern matching
haftmann
parents: 26011
diff changeset
   285
fun assemble_eqnss idx_of deps eqnss =
25944
haftmann
parents: 25935
diff changeset
   286
  let
haftmann
parents: 25935
diff changeset
   287
    fun prep_eqns (c, (vs, eqns)) =
25924
f974a1c64348 improved implementation
haftmann
parents: 25865
diff changeset
   288
      let
25944
haftmann
parents: 25935
diff changeset
   289
        val dicts = maps (fn (v, sort) => map_index (nbe_dict v o fst) sort) vs;
28337
93964076e7b8 case default fallback for NBE
haftmann
parents: 28296
diff changeset
   290
        val num_args = length dicts + ((length o fst o hd) eqns);
25944
haftmann
parents: 25935
diff changeset
   291
      in (c, (num_args, (dicts, eqns))) end;
haftmann
parents: 25935
diff changeset
   292
    val eqnss' = map prep_eqns eqnss;
haftmann
parents: 25935
diff changeset
   293
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   294
    fun assemble_constapp sym dss ts = 
25944
haftmann
parents: 25935
diff changeset
   295
      let
41118
b290841cd3b0 separated dictionary weakning into separate type
haftmann
parents: 41100
diff changeset
   296
        val ts' = (maps o map) assemble_dict dss @ ts;
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   297
      in case AList.lookup (op =) eqnss' sym
28337
93964076e7b8 case default fallback for NBE
haftmann
parents: 28296
diff changeset
   298
       of SOME (num_args, _) => if num_args <= length ts'
93964076e7b8 case default fallback for NBE
haftmann
parents: 28296
diff changeset
   299
            then let val (ts1, ts2) = chop num_args ts'
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   300
            in nbe_apps (nbe_apps_local idx_of 0 sym ts1) ts2
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   301
            end else nbe_apps (nbe_abss num_args (nbe_fun idx_of 0 sym)) ts'
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   302
        | NONE => if member (op =) deps sym
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   303
            then nbe_apps (nbe_fun idx_of 0 sym) ts'
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   304
            else nbe_apps_constr idx_of sym ts'
25924
f974a1c64348 improved implementation
haftmann
parents: 25865
diff changeset
   305
      end
41100
6c0940392fb4 dictionary constants must permit explicit weakening of classes;
haftmann
parents: 41068
diff changeset
   306
    and assemble_classrels classrels =
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   307
      fold_rev (fn classrel => assemble_constapp (Code_Symbol.Class_Relation classrel) [] o single) classrels
41118
b290841cd3b0 separated dictionary weakning into separate type
haftmann
parents: 41100
diff changeset
   308
    and assemble_dict (Dict (classrels, x)) =
b290841cd3b0 separated dictionary weakning into separate type
haftmann
parents: 41100
diff changeset
   309
          assemble_classrels classrels (assemble_plain_dict x)
b290841cd3b0 separated dictionary weakning into separate type
haftmann
parents: 41100
diff changeset
   310
    and assemble_plain_dict (Dict_Const (inst, dss)) =
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   311
          assemble_constapp (Code_Symbol.Class_Instance inst) dss []
41118
b290841cd3b0 separated dictionary weakning into separate type
haftmann
parents: 41100
diff changeset
   312
      | assemble_plain_dict (Dict_Var (v, (n, _))) =
b290841cd3b0 separated dictionary weakning into separate type
haftmann
parents: 41100
diff changeset
   313
          nbe_dict v n
25924
f974a1c64348 improved implementation
haftmann
parents: 25865
diff changeset
   314
28350
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   315
    fun assemble_iterm constapp =
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   316
      let
28350
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   317
        fun of_iterm match_cont t =
25944
haftmann
parents: 25935
diff changeset
   318
          let
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents: 27609
diff changeset
   319
            val (t', ts) = Code_Thingol.unfold_app t
28350
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   320
          in of_iapp match_cont t' (fold_rev (cons o of_iterm NONE) ts []) end
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   321
        and of_iapp match_cont (IConst { sym, dicts = dss, ... }) ts = constapp sym dss ts
31889
fb2c8a687529 all variable names are optional
haftmann
parents: 31888
diff changeset
   322
          | of_iapp match_cont (IVar v) ts = nbe_apps (nbe_bound_optional v) ts
31724
9b5a128cdb5c more appropriate syntax for IML abstraction
haftmann
parents: 31156
diff changeset
   323
          | of_iapp match_cont ((v, _) `|=> t) ts =
31888
626c075fd457 variable names in abstractions are optional
haftmann
parents: 31724
diff changeset
   324
              nbe_apps (nbe_abss 1 (ml_abs (ml_list [nbe_bound_optional v]) (of_iterm NONE t))) ts
48072
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 47609
diff changeset
   325
          | of_iapp match_cont (ICase { term = t, clauses = clauses, primitive = t0, ... }) ts =
28350
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   326
              nbe_apps (ml_cases (of_iterm NONE t)
48072
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 47609
diff changeset
   327
                (map (fn (p, t) => (of_iterm NONE p, of_iterm match_cont t)) clauses
28350
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   328
                  @ [("_", case match_cont of SOME s => s | NONE => of_iterm NONE t0)])) ts
25944
haftmann
parents: 25935
diff changeset
   329
      in of_iterm end;
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   330
28350
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   331
    fun subst_nonlin_vars args =
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   332
      let
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   333
        val vs = (fold o Code_Thingol.fold_varnames)
33002
f3f02f36a3e2 uniform use of Integer.add/mult/sum/prod;
wenzelm
parents: 32966
diff changeset
   334
          (fn v => AList.map_default (op =) (v, 0) (Integer.add 1)) args [];
28350
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   335
        val names = Name.make_context (map fst vs);
43329
84472e198515 tuned signature: Name.invent and Name.invent_names;
wenzelm
parents: 43323
diff changeset
   336
        fun declare v k ctxt =
84472e198515 tuned signature: Name.invent and Name.invent_names;
wenzelm
parents: 43323
diff changeset
   337
          let val vs = Name.invent ctxt v k
28350
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   338
          in (vs, fold Name.declare vs ctxt) end;
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   339
        val (vs_renames, _) = fold_map (fn (v, k) => if k > 1
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   340
          then declare v (k - 1) #>> (fn vs => (v, vs))
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   341
          else pair (v, [])) vs names;
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   342
        val samepairs = maps (fn (v, vs) => map (pair v) vs) vs_renames;
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   343
        fun subst_vars (t as IConst _) samepairs = (t, samepairs)
31889
fb2c8a687529 all variable names are optional
haftmann
parents: 31888
diff changeset
   344
          | subst_vars (t as IVar NONE) samepairs = (t, samepairs)
fb2c8a687529 all variable names are optional
haftmann
parents: 31888
diff changeset
   345
          | subst_vars (t as IVar (SOME v)) samepairs = (case AList.lookup (op =) samepairs v
fb2c8a687529 all variable names are optional
haftmann
parents: 31888
diff changeset
   346
             of SOME v' => (IVar (SOME v'), AList.delete (op =) v samepairs)
28350
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   347
              | NONE => (t, samepairs))
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   348
          | subst_vars (t1 `$ t2) samepairs = samepairs
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   349
              |> subst_vars t1
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   350
              ||>> subst_vars t2
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   351
              |>> (op `$)
48072
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 47609
diff changeset
   352
          | subst_vars (ICase { primitive = t, ... }) samepairs = subst_vars t samepairs;
28350
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   353
        val (args', _) = fold_map subst_vars args samepairs;
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   354
      in (samepairs, args') end;
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   355
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   356
    fun assemble_eqn sym dicts default_args (i, (args, rhs)) =
28337
93964076e7b8 case default fallback for NBE
haftmann
parents: 28296
diff changeset
   357
      let
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   358
        val match_cont = if Code_Symbol.is_value sym then NONE
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   359
          else SOME (nbe_apps_local idx_of (i + 1) sym (dicts @ default_args));
28350
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   360
        val assemble_arg = assemble_iterm
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   361
          (fn sym' => fn dss => fn ts => nbe_apps_constr idx_of sym' ((maps o map) (K "_") dss @ ts)) NONE;
35371
6c92eb394e3c tuned whitespace
haftmann
parents: 34251
diff changeset
   362
        val assemble_rhs = assemble_iterm assemble_constapp match_cont;
28350
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   363
        val (samepairs, args') = subst_nonlin_vars args;
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   364
        val s_args = map assemble_arg args';
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   365
        val s_rhs = if null samepairs then assemble_rhs rhs
41037
6d6f23b3a879 removed experimental equality checking of closures; acknowledge underapproximation of equality in function name
haftmann
parents: 40730
diff changeset
   366
          else ml_if (ml_and (map nbe_same samepairs))
55043
acefda71629b prefer indexes user for pattern matching to print concrete names for symbols, do not rely on printable unique identifiers
haftmann
parents: 54889
diff changeset
   367
            (assemble_rhs rhs) (the match_cont);
acefda71629b prefer indexes user for pattern matching to print concrete names for symbols, do not rely on printable unique identifiers
haftmann
parents: 54889
diff changeset
   368
        val eqns = case match_cont
acefda71629b prefer indexes user for pattern matching to print concrete names for symbols, do not rely on printable unique identifiers
haftmann
parents: 54889
diff changeset
   369
         of NONE => [([ml_list (rev (dicts @ s_args))], s_rhs)]
acefda71629b prefer indexes user for pattern matching to print concrete names for symbols, do not rely on printable unique identifiers
haftmann
parents: 54889
diff changeset
   370
          | SOME default_rhs =>
acefda71629b prefer indexes user for pattern matching to print concrete names for symbols, do not rely on printable unique identifiers
haftmann
parents: 54889
diff changeset
   371
              [([ml_list (rev (dicts @ map2 ml_as default_args s_args))], s_rhs),
acefda71629b prefer indexes user for pattern matching to print concrete names for symbols, do not rely on printable unique identifiers
haftmann
parents: 54889
diff changeset
   372
                ([ml_list (rev (dicts @ default_args))], default_rhs)]
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   373
      in (nbe_fun idx_of i sym, eqns) end;
28337
93964076e7b8 case default fallback for NBE
haftmann
parents: 28296
diff changeset
   374
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   375
    fun assemble_eqns (sym, (num_args, (dicts, eqns))) =
25944
haftmann
parents: 25935
diff changeset
   376
      let
28337
93964076e7b8 case default fallback for NBE
haftmann
parents: 28296
diff changeset
   377
        val default_args = map nbe_default
43329
84472e198515 tuned signature: Name.invent and Name.invent_names;
wenzelm
parents: 43323
diff changeset
   378
          (Name.invent Name.context "a" (num_args - length dicts));
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   379
        val eqns' = map_index (assemble_eqn sym dicts default_args) eqns
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   380
          @ (if Code_Symbol.is_value sym then [] else [(nbe_fun idx_of (length eqns) sym,
28337
93964076e7b8 case default fallback for NBE
haftmann
parents: 28296
diff changeset
   381
            [([ml_list (rev (dicts @ default_args))],
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   382
              nbe_apps_constr idx_of sym (dicts @ default_args))])]);
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   383
      in (eqns', nbe_abss num_args (nbe_fun idx_of 0 sym)) end;
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   384
25944
haftmann
parents: 25935
diff changeset
   385
    val (fun_vars, fun_vals) = map_split assemble_eqns eqnss';
55043
acefda71629b prefer indexes user for pattern matching to print concrete names for symbols, do not rely on printable unique identifiers
haftmann
parents: 54889
diff changeset
   386
    val deps_vars = ml_list (map (nbe_fun idx_of 0) deps);
28337
93964076e7b8 case default fallback for NBE
haftmann
parents: 28296
diff changeset
   387
  in ml_abs deps_vars (ml_Let (ml_fundefs (flat fun_vars)) (ml_list fun_vals)) end;
25944
haftmann
parents: 25935
diff changeset
   388
35371
6c92eb394e3c tuned whitespace
haftmann
parents: 34251
diff changeset
   389
39392
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   390
(* compile equations *)
25944
haftmann
parents: 25935
diff changeset
   391
39392
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   392
fun compile_eqnss thy nbe_program raw_deps [] = []
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   393
  | compile_eqnss thy nbe_program raw_deps eqnss =
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   394
      let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 41472
diff changeset
   395
        val ctxt = Proof_Context.init_global thy;
26064
65585de05a66 using integers for pattern matching
haftmann
parents: 26011
diff changeset
   396
        val (deps, deps_vals) = split_list (map_filter
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   397
          (fn dep => Option.map (fn univ => (dep, univ)) (fst ((Code_Symbol.Graph.get_node nbe_program dep)))) raw_deps);
26064
65585de05a66 using integers for pattern matching
haftmann
parents: 26011
diff changeset
   398
        val idx_of = raw_deps
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   399
          |> map (fn dep => (dep, snd (Code_Symbol.Graph.get_node nbe_program dep)))
26064
65585de05a66 using integers for pattern matching
haftmann
parents: 26011
diff changeset
   400
          |> AList.lookup (op =)
65585de05a66 using integers for pattern matching
haftmann
parents: 26011
diff changeset
   401
          |> (fn f => the o f);
65585de05a66 using integers for pattern matching
haftmann
parents: 26011
diff changeset
   402
        val s = assemble_eqnss idx_of deps eqnss;
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   403
        val cs = map fst eqnss;
25944
haftmann
parents: 25935
diff changeset
   404
      in
haftmann
parents: 25935
diff changeset
   405
        s
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   406
        |> traced (fn s => "\n--- code to be evaluated:\n" ^ s)
38931
5e84c11c4b8a evaluate takes ml context and ml expression parameter
haftmann
parents: 38676
diff changeset
   407
        |> pair ""
39911
2b4430847310 moved ML_Context.value to Code_Runtime
haftmann
parents: 39606
diff changeset
   408
        |> Code_Runtime.value ctxt univs_cookie
25944
haftmann
parents: 25935
diff changeset
   409
        |> (fn f => f deps_vals)
haftmann
parents: 25935
diff changeset
   410
        |> (fn univs => cs ~~ univs)
haftmann
parents: 25935
diff changeset
   411
      end;
25190
5cd8486c5a4f clarified implementation
haftmann
parents: 25167
diff changeset
   412
30672
beaadd5af500 more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents: 30288
diff changeset
   413
39392
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   414
(* extract equations from statements *)
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   415
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   416
fun dummy_const sym dss =
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   417
  IConst { sym = sym, typargs = [], dicts = dss,
48072
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 47609
diff changeset
   418
    dom = [], range = ITyVar "", annotate = false };
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 47609
diff changeset
   419
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   420
fun eqns_of_stmt (_, Code_Thingol.NoStmt) =
54889
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 52519
diff changeset
   421
      []
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   422
  | eqns_of_stmt (_, Code_Thingol.Fun ((_, []), _)) =
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   423
      []
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   424
  | eqns_of_stmt (sym_const, Code_Thingol.Fun (((vs, _), eqns), _)) =
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   425
      [(sym_const, (vs, map fst eqns))]
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents: 27609
diff changeset
   426
  | eqns_of_stmt (_, Code_Thingol.Datatypecons _) =
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   427
      []
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents: 27609
diff changeset
   428
  | eqns_of_stmt (_, Code_Thingol.Datatype _) =
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   429
      []
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   430
  | eqns_of_stmt (sym_class, Code_Thingol.Class (v, (classrels, classparams))) =
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   431
      let
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   432
        val syms = map Code_Symbol.Class_Relation classrels @ map (Code_Symbol.Constant o fst) classparams;
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   433
        val params = Name.invent Name.context "d" (length syms);
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   434
        fun mk (k, sym) =
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   435
          (sym, ([(v, [])],
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   436
            [([dummy_const sym_class [] `$$ map (IVar o SOME) params],
31889
fb2c8a687529 all variable names are optional
haftmann
parents: 31888
diff changeset
   437
              IVar (SOME (nth params k)))]));
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   438
      in map_index mk syms end
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents: 27609
diff changeset
   439
  | eqns_of_stmt (_, Code_Thingol.Classrel _) =
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   440
      []
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents: 27609
diff changeset
   441
  | eqns_of_stmt (_, Code_Thingol.Classparam _) =
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   442
      []
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   443
  | eqns_of_stmt (sym_inst, Code_Thingol.Classinst { class, tyco, vs, superinsts, inst_params, ... }) =
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   444
      [(sym_inst, (vs, [([], dummy_const (Code_Symbol.Type_Class class) [] `$$
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   445
        map (fn (class, dss) => dummy_const (Code_Symbol.Class_Instance (tyco, class)) dss) superinsts
52519
598addf65209 explicit hint for domain of class parameters in instance statements
haftmann
parents: 52474
diff changeset
   446
        @ map (IConst o fst o snd o fst) inst_params)]))];
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   447
39392
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   448
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   449
(* compile whole programs *)
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   450
39399
267235a14938 static nbe conversion
haftmann
parents: 39396
diff changeset
   451
fun ensure_const_idx name (nbe_program, (maxidx, idx_tab)) =
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   452
  if can (Code_Symbol.Graph.get_node nbe_program) name
39399
267235a14938 static nbe conversion
haftmann
parents: 39396
diff changeset
   453
  then (nbe_program, (maxidx, idx_tab))
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   454
  else (Code_Symbol.Graph.new_node (name, (NONE, maxidx)) nbe_program,
39399
267235a14938 static nbe conversion
haftmann
parents: 39396
diff changeset
   455
    (maxidx + 1, Inttab.update_new (maxidx, name) idx_tab));
267235a14938 static nbe conversion
haftmann
parents: 39396
diff changeset
   456
39392
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   457
fun compile_stmts thy stmts_deps =
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   458
  let
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   459
    val names = map (fst o fst) stmts_deps;
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   460
    val names_deps = map (fn ((name, _), deps) => (name, deps)) stmts_deps;
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   461
    val eqnss = maps (eqns_of_stmt o fst) stmts_deps;
26064
65585de05a66 using integers for pattern matching
haftmann
parents: 26011
diff changeset
   462
    val refl_deps = names_deps
25190
5cd8486c5a4f clarified implementation
haftmann
parents: 25167
diff changeset
   463
      |> maps snd
5cd8486c5a4f clarified implementation
haftmann
parents: 25167
diff changeset
   464
      |> distinct (op =)
26064
65585de05a66 using integers for pattern matching
haftmann
parents: 26011
diff changeset
   465
      |> fold (insert (op =)) names;
39392
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   466
    fun compile nbe_program = eqnss
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   467
      |> compile_eqnss thy nbe_program refl_deps
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   468
      |> rpair nbe_program;
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   469
  in
39399
267235a14938 static nbe conversion
haftmann
parents: 39396
diff changeset
   470
    fold ensure_const_idx refl_deps
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   471
    #> apfst (fold (fn (name, deps) => fold (curry Code_Symbol.Graph.add_edge name) deps) names_deps
26064
65585de05a66 using integers for pattern matching
haftmann
parents: 26011
diff changeset
   472
      #> compile
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   473
      #-> fold (fn (name, univ) => (Code_Symbol.Graph.map_node name o apfst) (K (SOME univ))))
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   474
  end;
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   475
39392
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   476
fun compile_program thy program =
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   477
  let
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   478
    fun add_stmts names (nbe_program, (maxidx, idx_tab)) = if exists ((can o Code_Symbol.Graph.get_node) nbe_program) names
39392
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   479
      then (nbe_program, (maxidx, idx_tab))
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   480
      else (nbe_program, (maxidx, idx_tab))
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   481
        |> compile_stmts thy (map (fn name => ((name, Code_Symbol.Graph.get_node program name),
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   482
          Code_Symbol.Graph.immediate_succs program name)) names);
28706
3fef773ae6b1 restored incremental code generation
haftmann
parents: 28663
diff changeset
   483
  in
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   484
    fold_rev add_stmts (Code_Symbol.Graph.strong_conn program)
28706
3fef773ae6b1 restored incremental code generation
haftmann
parents: 28663
diff changeset
   485
  end;
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   486
25944
haftmann
parents: 25935
diff changeset
   487
haftmann
parents: 25935
diff changeset
   488
(** evaluation **)
haftmann
parents: 25935
diff changeset
   489
39392
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   490
(* term evaluation by compilation *)
25944
haftmann
parents: 25935
diff changeset
   491
39392
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   492
fun compile_term thy nbe_program deps (vs : (string * sort) list, t) =
25924
f974a1c64348 improved implementation
haftmann
parents: 25865
diff changeset
   493
  let 
f974a1c64348 improved implementation
haftmann
parents: 25865
diff changeset
   494
    val dict_frees = maps (fn (v, sort) => map_index (curry DFree v o fst) sort) vs;
f974a1c64348 improved implementation
haftmann
parents: 25865
diff changeset
   495
  in
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   496
    (Code_Symbol.value, (vs, [([], t)]))
39392
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   497
    |> singleton (compile_eqnss thy nbe_program deps)
25924
f974a1c64348 improved implementation
haftmann
parents: 25865
diff changeset
   498
    |> snd
31064
ce37d8f48a9f treat frees driectly by the LCF kernel
haftmann
parents: 31049
diff changeset
   499
    |> (fn t => apps t (rev dict_frees))
25924
f974a1c64348 improved implementation
haftmann
parents: 25865
diff changeset
   500
  end;
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   501
35371
6c92eb394e3c tuned whitespace
haftmann
parents: 34251
diff changeset
   502
39392
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   503
(* reconstruction *)
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   504
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   505
fun typ_of_itype vs (tyco `%% itys) =
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   506
      Type (tyco, map (typ_of_itype vs) itys)
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   507
  | typ_of_itype vs (ITyVar v) =
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   508
      TFree ("'" ^ v, (the o AList.lookup (op =) vs) v);
30947
dd551284a300 re-engineering of evaluation conversions
haftmann
parents: 30942
diff changeset
   509
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   510
fun term_of_univ thy idx_tab t =
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   511
  let
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   512
    fun take_until f [] = []
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   513
      | take_until f (x :: xs) = if f x then [] else x :: take_until f xs;
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   514
    fun is_dict (Const (idx, _)) =
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   515
          (case Inttab.lookup idx_tab idx of
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   516
            SOME (Code_Symbol.Constant _) => false
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   517
          | _ => true)
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   518
      | is_dict (DFree _) = true
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   519
      | is_dict _ = false;
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   520
    fun const_of_idx idx =
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   521
      case Inttab.lookup idx_tab idx of SOME (Code_Symbol.Constant const) => const;
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   522
    fun of_apps bounds (t, ts) =
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   523
      fold_map (of_univ bounds) ts
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   524
      #>> (fn ts' => list_comb (t, rev ts'))
26064
65585de05a66 using integers for pattern matching
haftmann
parents: 26011
diff changeset
   525
    and of_univ bounds (Const (idx, ts)) typidx =
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   526
          let
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   527
            val ts' = take_until is_dict ts;
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   528
            val const = const_of_idx idx;
31957
a9742afd403e tuned interface of structure Code
haftmann
parents: 31889
diff changeset
   529
            val T = map_type_tvar (fn ((v, i), _) =>
37145
01aa36932739 renamed structure TypeInfer to Type_Infer, keeping the old name as legacy alias for some time;
wenzelm
parents: 36982
diff changeset
   530
              Type_Infer.param typidx (v ^ string_of_int i, []))
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   531
                (Sign.the_const_type thy const);
30022
1d8b8fa19074 maintain order of constructors in datatypes; clarified conventions for type schemes
haftmann
parents: 29272
diff changeset
   532
            val typidx' = typidx + 1;
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   533
          in of_apps bounds (Term.Const (const, T), ts') typidx' end
27499
150558266831 clarified code
haftmann
parents: 27264
diff changeset
   534
      | of_univ bounds (BVar (n, ts)) typidx =
150558266831 clarified code
haftmann
parents: 27264
diff changeset
   535
          of_apps bounds (Bound (bounds - n - 1), ts) typidx
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   536
      | of_univ bounds (t as Abs _) typidx =
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   537
          typidx
25924
f974a1c64348 improved implementation
haftmann
parents: 25865
diff changeset
   538
          |> of_univ (bounds + 1) (apps t [BVar (bounds, [])])
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   539
          |-> (fn t' => pair (Term.Abs ("u", dummyT, t')))
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   540
  in of_univ 0 t 0 |> fst end;
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   541
35371
6c92eb394e3c tuned whitespace
haftmann
parents: 34251
diff changeset
   542
39392
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   543
(* evaluation with type reconstruction *)
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   544
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   545
fun eval_term thy (nbe_program, idx_tab) ((vs0, (vs, ty)), t) deps =
39392
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   546
  let
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   547
    val ctxt = Syntax.init_pretty_global thy;
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   548
    val string_of_term = Syntax.string_of_term (Config.put show_types true ctxt);
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   549
    val ty' = typ_of_itype vs0 ty;
42405
13ecdb3057d8 split Type_Infer into early and late part, after Proof_Context;
wenzelm
parents: 42361
diff changeset
   550
    fun type_infer t =
13ecdb3057d8 split Type_Infer into early and late part, after Proof_Context;
wenzelm
parents: 42361
diff changeset
   551
      Syntax.check_term (Config.put Type_Infer_Context.const_sorts false ctxt)
13ecdb3057d8 split Type_Infer into early and late part, after Proof_Context;
wenzelm
parents: 42361
diff changeset
   552
        (Type.constraint ty' t);
39392
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   553
    fun check_tvars t =
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   554
      if null (Term.add_tvars t []) then t
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   555
      else error ("Illegal schematic type variables in normalized term: " ^ string_of_term t);
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   556
  in
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   557
    compile_term thy nbe_program deps (vs, t)
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   558
    |> term_of_univ thy idx_tab
39392
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   559
    |> traced (fn t => "Normalized:\n" ^ string_of_term t)
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   560
    |> type_infer
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   561
    |> traced (fn t => "Types inferred:\n" ^ string_of_term t)
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   562
    |> check_tvars
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   563
    |> traced (fn _ => "---\n")
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   564
  end;
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   565
39436
4a7d09da2b9c tuned whitespace
haftmann
parents: 39399
diff changeset
   566
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   567
(* function store *)
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   568
34173
458ced35abb8 reduced code generator cache to the baremost minimum
haftmann
parents: 33522
diff changeset
   569
structure Nbe_Functions = Code_Data
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   570
(
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   571
  type T = (Univ option * int) Code_Symbol.Graph.T * (int * Code_Symbol.symbol Inttab.table);
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   572
  val empty = (Code_Symbol.Graph.empty, (0, Inttab.empty));
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   573
);
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   574
39399
267235a14938 static nbe conversion
haftmann
parents: 39396
diff changeset
   575
fun compile ignore_cache thy program =
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   576
  let
39392
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   577
    val (nbe_program, (_, idx_tab)) =
39399
267235a14938 static nbe conversion
haftmann
parents: 39396
diff changeset
   578
      Nbe_Functions.change (if ignore_cache then NONE else SOME thy)
267235a14938 static nbe conversion
haftmann
parents: 39396
diff changeset
   579
        (compile_program thy program);
39392
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   580
  in (nbe_program, idx_tab) end;
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   581
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   582
47572
1e18bbfb40cb tuned heading
haftmann
parents: 44789
diff changeset
   583
(* evaluation oracle *)
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   584
30947
dd551284a300 re-engineering of evaluation conversions
haftmann
parents: 30942
diff changeset
   585
fun mk_equals thy lhs raw_rhs =
26747
f32fa5f5bdd1 moved 'trivial classes' to foundation of code generator
haftmann
parents: 26739
diff changeset
   586
  let
30947
dd551284a300 re-engineering of evaluation conversions
haftmann
parents: 30942
diff changeset
   587
    val ty = Thm.typ_of (Thm.ctyp_of_term lhs);
dd551284a300 re-engineering of evaluation conversions
haftmann
parents: 30942
diff changeset
   588
    val eq = Thm.cterm_of thy (Term.Const ("==", ty --> ty --> propT));
dd551284a300 re-engineering of evaluation conversions
haftmann
parents: 30942
diff changeset
   589
    val rhs = Thm.cterm_of thy raw_rhs;
dd551284a300 re-engineering of evaluation conversions
haftmann
parents: 30942
diff changeset
   590
  in Thm.mk_binop eq lhs rhs end;
26747
f32fa5f5bdd1 moved 'trivial classes' to foundation of code generator
haftmann
parents: 26739
diff changeset
   591
38669
9ff76d0f0610 refined and unified naming convention for dynamic code evaluation techniques
haftmann
parents: 37449
diff changeset
   592
val (_, raw_oracle) = Context.>>> (Context.map_theory_result
43619
3803869014aa proper @{binding} antiquotations (relevant for formal references);
wenzelm
parents: 43329
diff changeset
   593
  (Thm.add_oracle (@{binding normalization_by_evaluation},
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   594
    fn (thy, nbe_program_idx_tab, vsp_ty_t, deps, ct) =>
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   595
      mk_equals thy ct (eval_term thy nbe_program_idx_tab vsp_ty_t deps))));
31064
ce37d8f48a9f treat frees driectly by the LCF kernel
haftmann
parents: 31049
diff changeset
   596
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   597
fun oracle thy nbe_program_idx_tab vsp_ty_t deps ct =
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   598
  raw_oracle (thy, nbe_program_idx_tab, vsp_ty_t, deps, ct);
30942
1e246776f876 diagnostic commands now in code_thingol; tuned code of funny continuations
haftmann
parents: 30672
diff changeset
   599
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   600
fun dynamic_conv thy = lift_triv_classes_conv thy
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   601
  (Code_Thingol.dynamic_conv thy (oracle thy o compile false thy));
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   602
41184
5c6f44d22f51 simplified evaluation function names
haftmann
parents: 41118
diff changeset
   603
fun dynamic_value thy = lift_triv_classes_rew thy
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   604
  (Code_Thingol.dynamic_value thy I (eval_term thy o compile false thy));
39399
267235a14938 static nbe conversion
haftmann
parents: 39396
diff changeset
   605
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   606
fun static_conv thy consts = lift_triv_classes_conv thy
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   607
  (Code_Thingol.static_conv thy consts (K o oracle thy o compile true thy));
24839
199c48ec5a09 step towards proper purge operation
haftmann
parents: 24713
diff changeset
   608
41184
5c6f44d22f51 simplified evaluation function names
haftmann
parents: 41118
diff changeset
   609
fun static_value thy consts = lift_triv_classes_rew thy
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   610
  (Code_Thingol.static_value thy I consts (K o eval_term thy o compile true thy));
41184
5c6f44d22f51 simplified evaluation function names
haftmann
parents: 41118
diff changeset
   611
35371
6c92eb394e3c tuned whitespace
haftmann
parents: 34251
diff changeset
   612
39396
e9cad160aa0f dropped redundant normal_form command
haftmann
parents: 39392
diff changeset
   613
(** setup **)
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   614
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 41472
diff changeset
   615
val setup = Value.add_evaluator ("nbe", dynamic_value o Proof_Context.theory_of);
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   616
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   617
end;
28337
93964076e7b8 case default fallback for NBE
haftmann
parents: 28296
diff changeset
   618