src/Tools/nbe.ML
author nipkow
Thu, 17 Jul 2025 21:06:22 +0100
changeset 82885 5d2a599f88af
parent 82510 5601f5cce4c6
permissions -rw-r--r--
moved lemma
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
55757
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55167
diff changeset
     9
  val dynamic_conv: Proof.context -> conv
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55167
diff changeset
    10
  val dynamic_value: Proof.context -> term -> term
63164
72aaf69328fc optional timing for code generator conversions
haftmann
parents: 63162
diff changeset
    11
  val static_conv: { ctxt: Proof.context, consts: string list }
72aaf69328fc optional timing for code generator conversions
haftmann
parents: 63162
diff changeset
    12
    -> Proof.context -> conv
72aaf69328fc optional timing for code generator conversions
haftmann
parents: 63162
diff changeset
    13
  val static_value: { ctxt: Proof.context, consts: string list }
72aaf69328fc optional timing for code generator conversions
haftmann
parents: 63162
diff changeset
    14
    -> Proof.context -> term -> term
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
    15
82444
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
    16
  datatype Type =
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
    17
      Type of string * Type list
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
    18
    | TParam of string
25204
36cf92f63a44 replaced Secure.evaluate by ML_Context.evaluate;
wenzelm
parents: 25190
diff changeset
    19
  datatype Univ =
82444
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
    20
      Const of (int * Type list) * Univ list (*named (uninterpreted) constants*)
25924
f974a1c64348 improved implementation
haftmann
parents: 25865
diff changeset
    21
    | DFree of string * int                  (*free (uninterpreted) dictionary parameters*)
82444
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
    22
    | BVar of int * Univ list                (*bound variables, named*)
28350
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
    23
    | Abs of (int * (Univ list -> Univ)) * Univ list
25924
f974a1c64348 improved implementation
haftmann
parents: 25865
diff changeset
    24
  val apps: Univ -> Univ list -> Univ        (*explicit applications*)
25944
haftmann
parents: 25935
diff changeset
    25
  val abss: int -> (Univ list -> Univ) -> Univ
28337
93964076e7b8 case default fallback for NBE
haftmann
parents: 28296
diff changeset
    26
                                             (*abstractions as closures*)
41037
6d6f23b3a879 removed experimental equality checking of closures; acknowledge underapproximation of equality in function name
haftmann
parents: 40730
diff changeset
    27
  val same: Univ * Univ -> bool
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
    28
82444
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
    29
  val put_result: (unit -> (Type list -> Univ) list -> (Type list -> Univ) list)
63164
72aaf69328fc optional timing for code generator conversions
haftmann
parents: 63162
diff changeset
    30
    -> Proof.context -> Proof.context
57435
312660c1a70a modernized diagnostic options
haftmann
parents: 57174
diff changeset
    31
  val trace: bool Config.T
25924
f974a1c64348 improved implementation
haftmann
parents: 25865
diff changeset
    32
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    33
  val add_const_alias: thm -> theory -> theory
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
    34
end;
d86867645f4f nbe improved
haftmann
parents:
diff changeset
    35
d86867645f4f nbe improved
haftmann
parents:
diff changeset
    36
structure Nbe: NBE =
d86867645f4f nbe improved
haftmann
parents:
diff changeset
    37
struct
d86867645f4f nbe improved
haftmann
parents:
diff changeset
    38
d86867645f4f nbe improved
haftmann
parents:
diff changeset
    39
(* generic non-sense *)
d86867645f4f nbe improved
haftmann
parents:
diff changeset
    40
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 63164
diff changeset
    41
val trace = Attrib.setup_config_bool \<^binding>\<open>nbe_trace\<close> (K false);
57435
312660c1a70a modernized diagnostic options
haftmann
parents: 57174
diff changeset
    42
fun traced ctxt f x = if Config.get ctxt trace then (tracing (f x); x) else x;
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
    43
d86867645f4f nbe improved
haftmann
parents:
diff changeset
    44
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    45
(** certificates and oracle for "trivial type classes" **)
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    46
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33002
diff changeset
    47
structure Triv_Class_Data = Theory_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
  type T = (class * thm) list;
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    50
  val empty = [];
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33002
diff changeset
    51
  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
    52
);
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    53
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    54
fun add_const_alias thm thy =
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    55
  let
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    56
    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
    57
     of SOME ofclass_eq => ofclass_eq
61268
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 61096
diff changeset
    58
      | _ => error ("Bad certificate: " ^ Thm.string_of_thm_global thy thm);
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    59
    val (T, class) = case try Logic.dest_of_class ofclass
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    60
     of SOME T_class => T_class
61268
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 61096
diff changeset
    61
      | _ => error ("Bad certificate: " ^ Thm.string_of_thm_global thy thm);
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    62
    val tvar = case try Term.dest_TVar T
51685
385ef6706252 more standard module name Axclass (according to file name);
wenzelm
parents: 48072
diff changeset
    63
     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
    64
          then tvar
61268
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 61096
diff changeset
    65
          else error ("Bad sort: " ^ Thm.string_of_thm_global thy thm)
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 61096
diff changeset
    66
      | _ => error ("Bad type: " ^ Thm.string_of_thm_global thy thm);
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    67
    val _ = if Term.add_tvars eqn [] = [tvar] then ()
61268
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 61096
diff changeset
    68
      else error ("Inconsistent type: " ^ Thm.string_of_thm_global thy thm);
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    69
    val lhs_rhs = case try Logic.dest_equals eqn
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    70
     of SOME lhs_rhs => lhs_rhs
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    71
      | _ => error ("Not an equation: " ^ Syntax.string_of_term_global thy eqn);
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58397
diff changeset
    72
    val c_c' = case try (apply2 (Axclass.unoverload_const thy o dest_Const)) lhs_rhs
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    73
     of SOME c_c' => c_c'
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    74
      | _ => error ("Not an equation with two constants: "
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    75
          ^ Syntax.string_of_term_global thy eqn);
51685
385ef6706252 more standard module name Axclass (according to file name);
wenzelm
parents: 48072
diff changeset
    76
    val _ = if the_list (Axclass.class_of_param thy (snd c_c')) = [class] then ()
61268
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 61096
diff changeset
    77
      else error ("Inconsistent class: " ^ Thm.string_of_thm_global thy thm);
61096
cc5f4b8ac05b trim context for persistent storage;
wenzelm
parents: 60642
diff changeset
    78
  in Triv_Class_Data.map (AList.update (op =) (class, Thm.trim_context thm)) thy end;
32544
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
local
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    81
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    82
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
    83
78795
f7e972d567f3 clarified signature: more concise variations on implicit theory setup;
wenzelm
parents: 77707
diff changeset
    84
val (_, triv_of_class) =
f7e972d567f3 clarified signature: more concise variations on implicit theory setup;
wenzelm
parents: 77707
diff changeset
    85
  Theory.setup_result (Thm.add_oracle (\<^binding>\<open>triv_of_class\<close>,
f7e972d567f3 clarified signature: more concise variations on implicit theory setup;
wenzelm
parents: 77707
diff changeset
    86
    fn (thy, T, class) => Thm.global_cterm_of thy (Logic.mk_of_class (T, class))));
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    87
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    88
in
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    89
63158
534f16b0ca39 explicit quasi-global context for nbe conversions -- works around quasi-global type variable handling in lift_triv_classes_conv
haftmann
parents: 63157
diff changeset
    90
fun lift_triv_classes_conv orig_ctxt conv ct =
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    91
  let
63158
534f16b0ca39 explicit quasi-global context for nbe conversions -- works around quasi-global type variable handling in lift_triv_classes_conv
haftmann
parents: 63157
diff changeset
    92
    val thy = Proof_Context.theory_of orig_ctxt;
534f16b0ca39 explicit quasi-global context for nbe conversions -- works around quasi-global type variable handling in lift_triv_classes_conv
haftmann
parents: 63157
diff changeset
    93
    val ctxt = Proof_Context.init_global thy;
534f16b0ca39 explicit quasi-global context for nbe conversions -- works around quasi-global type variable handling in lift_triv_classes_conv
haftmann
parents: 63157
diff changeset
    94
      (*FIXME quasi-global context*)
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    95
    val algebra = Sign.classes_of thy;
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
    96
    val triv_classes = get_triv_classes thy;
36982
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
    97
    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
    98
    fun mk_entry (v, sort) =
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
    99
      let
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   100
        val T = TFree (v, sort);
60322
05fabeb0130a clarified context;
wenzelm
parents: 59642
diff changeset
   101
        val cT = Thm.ctyp_of ctxt T;
36982
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   102
        val triv_sort = additional_classes sort;
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   103
      in
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   104
        (v, (Sorts.inter_sort algebra (sort, triv_sort),
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   105
          (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
   106
            @ 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
   107
      end;
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   108
    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
   109
    fun instantiate thm =
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   110
      let
60642
48dd1cefb4ae simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents: 60322
diff changeset
   111
        val tvars =
48dd1cefb4ae simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents: 60322
diff changeset
   112
          Term.add_tvars (#1 (Logic.dest_equals (Logic.strip_imp_concl (Thm.prop_of thm)))) [];
48dd1cefb4ae simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents: 60322
diff changeset
   113
        val instT = map2 (fn v => fn (_, (_, (cT, _))) => (v, cT)) tvars vs_tab;
74282
c2ee8d993d6a clarified signature: more scalable operations;
wenzelm
parents: 69593
diff changeset
   114
      in Thm.instantiate (TVars.make instT, Vars.empty) thm end;
36982
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   115
    fun of_class (TFree (v, _), class) =
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   116
          the (AList.lookup (op =) ((snd o snd o the o AList.lookup (op =) vs_tab) v) class)
60322
05fabeb0130a clarified context;
wenzelm
parents: 59642
diff changeset
   117
      | of_class (T, _) = error ("Bad type " ^ Syntax.string_of_typ ctxt T);
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   118
    fun strip_of_class thm =
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   119
      let
36982
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   120
        val prems_of_class = Thm.prop_of thm
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   121
          |> Logic.strip_imp_prems
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   122
          |> map (Logic.dest_of_class #> of_class);
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   123
      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
   124
  in
c3a04614c710 reactivated Thm.legacy_unconstrainT for Nbe.lift_triv_classes_conv;
wenzelm
parents: 36768
diff changeset
   125
    ct
60322
05fabeb0130a clarified context;
wenzelm
parents: 59642
diff changeset
   126
    |> Thm.term_of
05fabeb0130a clarified context;
wenzelm
parents: 59642
diff changeset
   127
    |> (map_types o map_type_tfree)
47609
b3dab1892cda dropped dead code
haftmann
parents: 47573
diff changeset
   128
        (fn (v, _) => TFree (v, (fst o the o AList.lookup (op =) vs_tab) v))
60322
05fabeb0130a clarified context;
wenzelm
parents: 59642
diff changeset
   129
    |> Thm.cterm_of ctxt
63158
534f16b0ca39 explicit quasi-global context for nbe conversions -- works around quasi-global type variable handling in lift_triv_classes_conv
haftmann
parents: 63157
diff changeset
   130
    |> conv ctxt
36982
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   131
    |> 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
   132
    |> Thm.varifyT_global
36982
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   133
    |> Thm.unconstrainT
1d4478a797c2 new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents: 36960
diff changeset
   134
    |> instantiate
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   135
    |> strip_of_class
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   136
  end;
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   137
55757
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55167
diff changeset
   138
fun lift_triv_classes_rew ctxt rew t =
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   139
  let
55757
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55167
diff changeset
   140
    val thy = Proof_Context.theory_of ctxt;
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   141
    val algebra = Sign.classes_of thy;
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   142
    val triv_classes = get_triv_classes thy;
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   143
    val vs = Term.add_tfrees t [];
52473
a2407f62a29f do not choke on type variables emerging during rewriting
haftmann
parents: 51685
diff changeset
   144
  in
a2407f62a29f do not choke on type variables emerging during rewriting
haftmann
parents: 51685
diff changeset
   145
    t
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   146
    |> (map_types o map_type_tfree)
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   147
        (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
   148
    |> rew
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   149
    |> (map_types o map_type_tfree)
52473
a2407f62a29f do not choke on type variables emerging during rewriting
haftmann
parents: 51685
diff changeset
   150
        (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
   151
  end;
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   152
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   153
end;
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   154
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   155
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   156
(** the semantic universe **)
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   157
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   158
(*
82373
2819f79792b9 spelling
haftmann
parents: 81513
diff changeset
   159
   Functions are given by their semantic function value. To avoid
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   160
   trouble with the ML-type system, these functions have the most
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   161
   generic type, that is "Univ list -> Univ". The calling convention is
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   162
   that the arguments come as a list, the last argument first. In
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   163
   other words, a function call that usually would look like
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   164
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   165
   f x_1 x_2 ... x_n   or   f(x_1,x_2, ..., x_n)
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   166
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   167
   would be in our convention called as
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   168
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   169
              f [x_n,..,x_2,x_1]
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   170
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   171
   Moreover, to handle functions that are still waiting for some
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   172
   arguments we have additionally a list of arguments collected to far
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   173
   and the number of arguments we're still waiting for.
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   174
*)
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   175
82444
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   176
datatype Type =
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   177
    Type of string * Type list
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   178
  | TParam of string
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   179
25204
36cf92f63a44 replaced Secure.evaluate by ML_Context.evaluate;
wenzelm
parents: 25190
diff changeset
   180
datatype Univ =
82444
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   181
    Const of (int * Type list) * Univ list (*named (uninterpreted) constants*)
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   182
  | DFree of string * int                  (*free (uninterpreted) dictionary parameters*)
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   183
  | BVar of int * Univ list                (*bound variables, named*)
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   184
  | Abs of (int * (Univ list -> Univ)) * Univ list
82444
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   185
                                           (*abstractions as closures*)
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   186
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   187
(* constructor functions *)
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   188
25944
haftmann
parents: 25935
diff changeset
   189
fun abss n f = Abs ((n, f), []);
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   190
fun apps (Abs ((n, f), us)) ws = let val k = n - length ws in
27499
150558266831 clarified code
haftmann
parents: 27264
diff changeset
   191
      case int_ord (k, 0)
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   192
       of EQUAL => f (ws @ us)
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   193
        | LESS => let val (ws1, ws2) = chop (~ k) ws in apps (f (ws2 @ us)) ws1 end
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   194
        | GREATER => Abs ((k, f), ws @ us) (*note: reverse convention also for apps!*)
27499
150558266831 clarified code
haftmann
parents: 27264
diff changeset
   195
      end
25924
f974a1c64348 improved implementation
haftmann
parents: 25865
diff changeset
   196
  | apps (Const (name, xs)) ys = Const (name, ys @ xs)
27499
150558266831 clarified code
haftmann
parents: 27264
diff changeset
   197
  | apps (BVar (n, xs)) ys = BVar (n, ys @ xs);
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   198
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   199
fun same_type (Type (tyco1, ty1), Type (tyco2, tys2)) =
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   200
      (tyco1 = tyco2) andalso eq_list same_type (ty1, tys2)
82444
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   201
  | same_type (TParam v1, TParam v2) = (v1 = v2)
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   202
  | same_type _ = false;
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   203
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   204
fun same (Const ((k1, typargs1), us1), Const ((k2, typargs2), us2)) =
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   205
      (k1 = k2) andalso eq_list same_type (typargs1, typargs2) andalso eq_list same (us1, us2)
82444
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   206
  | same (DFree (n1, i1), DFree (n2, i2)) = (n1 = n2) andalso (i1 = i2)
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   207
  | same (BVar (i1, us1), BVar (i2, us2)) = (i1 = i2) andalso eq_list same (us1, us2)
41037
6d6f23b3a879 removed experimental equality checking of closures; acknowledge underapproximation of equality in function name
haftmann
parents: 40730
diff changeset
   208
  | same _ = false;
40730
2aa0390a2da7 nbe decides equality of abstractions by extensionality
haftmann
parents: 40564
diff changeset
   209
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   210
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   211
(** assembling and compiling ML code from terms **)
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   212
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   213
(* abstract ML syntax *)
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   214
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   215
infix 9 `$` `$$`;
82453
haftmann
parents: 82447
diff changeset
   216
infix 8 `*`;
82510
haftmann
parents: 82509
diff changeset
   217
82453
haftmann
parents: 82447
diff changeset
   218
fun e1 `$` e2 = enclose "(" ")" (e1 ^ " " ^ e2);
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   219
fun e `$$` [] = e
82453
haftmann
parents: 82447
diff changeset
   220
  | e `$$` es = enclose "(" ")" (e ^ " " ^ implode_space es);
haftmann
parents: 82447
diff changeset
   221
fun ml_abs v e = enclose "(" ")" ("fn " ^ v ^ " => " ^ e);
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   222
82453
haftmann
parents: 82447
diff changeset
   223
fun ml_cases e cs = enclose "(" ")"
haftmann
parents: 82447
diff changeset
   224
  ("case " ^ e ^ " of " ^ space_implode " | " (map (fn (p, e) => p ^ " => " ^ e) cs));
82506
289b18955960 revamped generation of functions
haftmann
parents: 82454
diff changeset
   225
fun ml_let d e = "\n  let\n" ^ prefix_lines "    " d ^ "\n  in " ^ e ^ " end";
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   226
28350
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   227
fun ml_and [] = "true"
82453
haftmann
parents: 82447
diff changeset
   228
  | ml_and [e] = e
haftmann
parents: 82447
diff changeset
   229
  | ml_and es = enclose "(" ")" (space_implode " andalso " es);
82506
289b18955960 revamped generation of functions
haftmann
parents: 82454
diff changeset
   230
fun ml_if b e1 e2 = enclose "(" ")" (implode_space ["if", b, "then", e1, "else", e2]);
28350
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   231
82453
haftmann
parents: 82447
diff changeset
   232
fun e1 `*` e2 = enclose "(" ")" (e1 ^ ", " ^ e2);
haftmann
parents: 82447
diff changeset
   233
fun ml_list es = enclose "[" "]" (commas es);
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   234
82506
289b18955960 revamped generation of functions
haftmann
parents: 82454
diff changeset
   235
fun ml_exc s = enclose "(" ")" ("raise Fail " ^ quote s);
289b18955960 revamped generation of functions
haftmann
parents: 82454
diff changeset
   236
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   237
fun ml_fundefs ([(name, [([], e)])]) =
82506
289b18955960 revamped generation of functions
haftmann
parents: 82454
diff changeset
   238
      "val " ^ name ^ " = " ^ e ^ ""
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   239
  | ml_fundefs (eqs :: eqss) =
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   240
      let
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   241
        fun fundef (name, eqs) =
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   242
          let
80910
406a85a25189 clarified signature: more explicit operations;
wenzelm
parents: 78795
diff changeset
   243
            fun eqn (es, e) = name ^ " " ^ implode_space es ^ " = " ^ e
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   244
          in space_implode "\n  | " (map eqn eqs) end;
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   245
      in
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   246
        (prefix "fun " o fundef) eqs :: map (prefix "and " o fundef) eqss
26931
aa226d8405a8 cat_lines;
wenzelm
parents: 26747
diff changeset
   247
        |> cat_lines
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   248
      end;
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   249
35371
6c92eb394e3c tuned whitespace
haftmann
parents: 34251
diff changeset
   250
25944
haftmann
parents: 25935
diff changeset
   251
(* nbe specific syntax and sandbox communication *)
haftmann
parents: 25935
diff changeset
   252
41472
f6ab14e61604 misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents: 41346
diff changeset
   253
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
   254
(
82444
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   255
  type T = unit -> (Type list -> Univ) list -> (Type list -> Univ) list;
59153
wenzelm
parents: 59151
diff changeset
   256
  val empty: T = fn () => raise Fail "Univs";
wenzelm
parents: 59151
diff changeset
   257
  fun init _ = empty;
39388
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39290
diff changeset
   258
);
59153
wenzelm
parents: 59151
diff changeset
   259
val get_result = Univs.get;
39388
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39290
diff changeset
   260
val put_result = Univs.put;
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   261
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   262
local
59151
a012574b78e7 proper exception for internal program failure, not user-error;
wenzelm
parents: 59058
diff changeset
   263
  val prefix = "Nbe.";
a012574b78e7 proper exception for internal program failure, not user-error;
wenzelm
parents: 59058
diff changeset
   264
  val name_put = prefix ^ "put_result";
41068
7e643e07be7f tuned whitespace
haftmann
parents: 41037
diff changeset
   265
  val name_const = prefix ^ "Const";
82444
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   266
  val name_type = prefix ^ "Type";
59151
a012574b78e7 proper exception for internal program failure, not user-error;
wenzelm
parents: 59058
diff changeset
   267
  val name_abss = prefix ^ "abss";
a012574b78e7 proper exception for internal program failure, not user-error;
wenzelm
parents: 59058
diff changeset
   268
  val name_apps = prefix ^ "apps";
a012574b78e7 proper exception for internal program failure, not user-error;
wenzelm
parents: 59058
diff changeset
   269
  val name_same = prefix ^ "same";
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   270
in
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   271
59153
wenzelm
parents: 59151
diff changeset
   272
val univs_cookie = (get_result, put_result, name_put);
25944
haftmann
parents: 25935
diff changeset
   273
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   274
(*
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   275
  Convention: parameters representing ("assembled") string representations of logical entities
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   276
  are prefixed with an "a_" -- unless they are an unqualified name ready to become part of
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   277
  an ML identifier.
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   278
*)
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   279
82444
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   280
fun nbe_tparam v = "t_" ^ v;
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   281
fun nbe_dict v n = "d_" ^ v ^ "_" ^ string_of_int n;
82506
289b18955960 revamped generation of functions
haftmann
parents: 82454
diff changeset
   282
fun nbe_global_param v = "w_" ^ v;
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   283
fun nbe_bound v = "v_" ^ v;
31888
626c075fd457 variable names in abstractions are optional
haftmann
parents: 31724
diff changeset
   284
fun nbe_bound_optional NONE = "_"
35500
118cda173627 dropped superfluous naming
haftmann
parents: 35371
diff changeset
   285
  | nbe_bound_optional (SOME v) = nbe_bound v;
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   286
fun nbe_type a_sym a_tys = name_type `$` (quote a_sym `*` ml_list a_tys);
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   287
fun nbe_fun a_sym a_typargs = a_sym `$` ml_list a_typargs;
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   288
82506
289b18955960 revamped generation of functions
haftmann
parents: 82454
diff changeset
   289
(*note: these are the "turning spots" where proper argument order is established!*)
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   290
fun nbe_apps a_u [] = a_u
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   291
  | nbe_apps a_u a_us = name_apps `$$` [a_u, ml_list (rev a_us)];
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   292
fun nbe_apps_fun a_sym a_typargs a_us = nbe_fun a_sym a_typargs `$` ml_list (rev a_us);
82508
haftmann
parents: 82507
diff changeset
   293
fun nbe_apps_fallback_fun a_sym a_us = a_sym `$$` a_us;
82510
haftmann
parents: 82509
diff changeset
   294
fun nbe_apps_const a_sym a_typargs a_us = name_const `$` ((a_sym `*` ml_list a_typargs) `*` ml_list (rev a_us));
haftmann
parents: 82509
diff changeset
   295
fun nbe_apps_constpat a_sym a_us = name_const `$` ((a_sym `*` "_") `*` ml_list (rev a_us));
25924
f974a1c64348 improved implementation
haftmann
parents: 25865
diff changeset
   296
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   297
fun nbe_abss 0 f = f `$` ml_list []
25944
haftmann
parents: 25935
diff changeset
   298
  | nbe_abss n f = name_abss `$$` [string_of_int n, f];
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   299
82453
haftmann
parents: 82447
diff changeset
   300
fun nbe_same (v1, v2) = enclose "(" ")" (name_same `$` (nbe_bound v1 `*` nbe_bound v2));
28350
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   301
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   302
end;
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   303
55150
0940309ed8f1 less clumsy namespace
haftmann
parents: 55147
diff changeset
   304
open Basic_Code_Symbol;
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents: 27609
diff changeset
   305
open Basic_Code_Thingol;
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   306
35371
6c92eb394e3c tuned whitespace
haftmann
parents: 34251
diff changeset
   307
25865
a141d6bfd398 tuned comment
haftmann
parents: 25204
diff changeset
   308
(* code generation *)
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   309
82375
haftmann
parents: 82374
diff changeset
   310
fun subst_nonlin_vars args =
25944
haftmann
parents: 25935
diff changeset
   311
  let
82375
haftmann
parents: 82374
diff changeset
   312
    val vs = (fold o Code_Thingol.fold_varnames)
haftmann
parents: 82374
diff changeset
   313
      (fn v => AList.map_default (op =) (v, 0) (Integer.add 1)) args [];
haftmann
parents: 82374
diff changeset
   314
    val names = Name.make_context (map fst vs);
haftmann
parents: 82374
diff changeset
   315
    val (vs_renames, _) = fold_map (fn (v, k) => if k > 1
haftmann
parents: 82374
diff changeset
   316
      then Name.invent' v (k - 1) #>> (fn vs => (v, vs))
haftmann
parents: 82374
diff changeset
   317
      else pair (v, [])) vs names;
haftmann
parents: 82374
diff changeset
   318
    val samepairs = maps (fn (v, vs) => map (pair v) vs) vs_renames;
haftmann
parents: 82374
diff changeset
   319
    fun subst_vars (t as IConst _) samepairs = (t, samepairs)
haftmann
parents: 82374
diff changeset
   320
      | subst_vars (t as IVar NONE) samepairs = (t, samepairs)
haftmann
parents: 82374
diff changeset
   321
      | subst_vars (t as IVar (SOME v)) samepairs = (case AList.lookup (op =) samepairs v
haftmann
parents: 82374
diff changeset
   322
         of SOME v' => (IVar (SOME v'), AList.delete (op =) v samepairs)
haftmann
parents: 82374
diff changeset
   323
          | NONE => (t, samepairs))
haftmann
parents: 82374
diff changeset
   324
      | subst_vars (t1 `$ t2) samepairs = samepairs
haftmann
parents: 82374
diff changeset
   325
          |> subst_vars t1
haftmann
parents: 82374
diff changeset
   326
          ||>> subst_vars t2
haftmann
parents: 82374
diff changeset
   327
          |>> (op `$)
haftmann
parents: 82374
diff changeset
   328
      | subst_vars (ICase { primitive = t, ... }) samepairs = subst_vars t samepairs;
haftmann
parents: 82374
diff changeset
   329
    val (args', _) = fold_map subst_vars args samepairs;
haftmann
parents: 82374
diff changeset
   330
  in (samepairs, args') end;
25944
haftmann
parents: 25935
diff changeset
   331
82375
haftmann
parents: 82374
diff changeset
   332
fun preprocess_eqns (sym, (vs, eqns)) =
haftmann
parents: 82374
diff changeset
   333
  let
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   334
    val a_tparams = map (fn (v, _) => nbe_tparam v) vs;
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   335
    val a_dict_params = maps (fn (v, sort) => map_index (nbe_dict v o fst) sort) vs;
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   336
    val num_args = length a_dict_params + ((length o fst o hd) eqns);
82506
289b18955960 revamped generation of functions
haftmann
parents: 82454
diff changeset
   337
    val a_global_params = map nbe_global_param (Name.invent_global "a" (num_args - length a_dict_params));
289b18955960 revamped generation of functions
haftmann
parents: 82454
diff changeset
   338
  in (sym, (num_args, (a_tparams, a_dict_params, a_global_params, (map o apfst) subst_nonlin_vars eqns))) end;
82444
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   339
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   340
fun assemble_type (tyco `%% tys) = nbe_type tyco (map assemble_type tys)
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   341
  | assemble_type (ITyVar v) = nbe_tparam v
82375
haftmann
parents: 82374
diff changeset
   342
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   343
fun assemble_preprocessed_eqnss ctxt idx_of_sym deps eqnss =
82375
haftmann
parents: 82374
diff changeset
   344
  let
82507
haftmann
parents: 82506
diff changeset
   345
    fun suffixed_fun_ident suffix sym = space_implode "_"
82510
haftmann
parents: 82509
diff changeset
   346
      ["c", if Code_Symbol.is_value sym then "0" else string_of_int (idx_of_sym sym),
82507
haftmann
parents: 82506
diff changeset
   347
      Code_Symbol.default_base sym, suffix];
haftmann
parents: 82506
diff changeset
   348
    val fun_ident = suffixed_fun_ident "nbe";
82508
haftmann
parents: 82507
diff changeset
   349
    fun fallback_fun_ident i = suffixed_fun_ident (string_of_int i);
82510
haftmann
parents: 82509
diff changeset
   350
    fun const_ident sym =
82374
haftmann
parents: 82373
diff changeset
   351
      if Config.get ctxt trace
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   352
      then string_of_int (idx_of_sym sym) ^ " (*" ^ Code_Symbol.default_base sym ^ "*)"
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   353
      else string_of_int (idx_of_sym sym);
82374
haftmann
parents: 82373
diff changeset
   354
82506
289b18955960 revamped generation of functions
haftmann
parents: 82454
diff changeset
   355
    fun assemble_fun sym = nbe_fun (fun_ident sym);
289b18955960 revamped generation of functions
haftmann
parents: 82454
diff changeset
   356
    fun assemble_app_fun sym = nbe_apps_fun (fun_ident sym);
82508
haftmann
parents: 82507
diff changeset
   357
    fun assemble_app_fallback_fun i sym = nbe_apps_fallback_fun (fallback_fun_ident i sym);
82510
haftmann
parents: 82509
diff changeset
   358
    fun assemble_app_const sym = nbe_apps_const (const_ident sym);
haftmann
parents: 82509
diff changeset
   359
    fun assemble_app_constpat sym = nbe_apps_constpat (const_ident sym);
82374
haftmann
parents: 82373
diff changeset
   360
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   361
    fun assemble_constapp sym typargs dictss a_ts =
25944
haftmann
parents: 25935
diff changeset
   362
      let
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   363
        val a_typargs = map (assemble_type) typargs;
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   364
        val a_ts' = (maps o map) assemble_dict (map2 (fn ty => map (fn dict => (ty, dict))) typargs dictss) @ a_ts;
82375
haftmann
parents: 82374
diff changeset
   365
      in case AList.lookup (op =) eqnss sym
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   366
       of SOME (num_args, _) => if num_args <= length a_ts'
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   367
            then let val (a_ts1, a_ts2) = chop num_args a_ts'
82506
289b18955960 revamped generation of functions
haftmann
parents: 82454
diff changeset
   368
            in nbe_apps (assemble_app_fun sym a_typargs a_ts1) a_ts2
289b18955960 revamped generation of functions
haftmann
parents: 82454
diff changeset
   369
            end else nbe_apps (nbe_abss num_args (assemble_fun sym a_typargs)) a_ts'
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   370
        | NONE => if member (op =) deps sym
82506
289b18955960 revamped generation of functions
haftmann
parents: 82454
diff changeset
   371
            then nbe_apps (assemble_fun sym a_typargs) a_ts'
82510
haftmann
parents: 82509
diff changeset
   372
            else assemble_app_const sym a_typargs a_ts'
25924
f974a1c64348 improved implementation
haftmann
parents: 25865
diff changeset
   373
      end
41100
6c0940392fb4 dictionary constants must permit explicit weakening of classes;
haftmann
parents: 41068
diff changeset
   374
    and assemble_classrels classrels =
82444
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   375
      fold_rev (fn classrel => assemble_constapp (Class_Relation classrel) [] [] o single) classrels
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   376
    and assemble_dict (ty, Dict (classrels, dict)) =
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   377
          assemble_classrels classrels (assemble_plain_dict ty dict)
82447
741f6f6df144 reflect nested lists in variables names
haftmann
parents: 82444
diff changeset
   378
    and assemble_plain_dict (_ `%% tys) (Dict_Const (inst, dictss)) =
741f6f6df144 reflect nested lists in variables names
haftmann
parents: 82444
diff changeset
   379
          assemble_constapp (Class_Instance inst) tys (map snd dictss) []
82444
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   380
      | assemble_plain_dict _ (Dict_Var { var, index, ... }) =
82509
haftmann
parents: 82508
diff changeset
   381
          nbe_dict var index;
25924
f974a1c64348 improved implementation
haftmann
parents: 25865
diff changeset
   382
82509
haftmann
parents: 82508
diff changeset
   383
    fun assemble_iterm is_pat a_match_fallback t =
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   384
      let
82509
haftmann
parents: 82508
diff changeset
   385
        fun assemble_app (IConst { sym, typargs, dictss, ... }) =
82510
haftmann
parents: 82509
diff changeset
   386
              if is_pat then fn a_ts => assemble_app_constpat sym ((maps o map) (K "_") dictss @ a_ts)
82509
haftmann
parents: 82508
diff changeset
   387
              else assemble_constapp sym typargs dictss
haftmann
parents: 82508
diff changeset
   388
          | assemble_app (IVar v) = nbe_apps (nbe_bound_optional v)
haftmann
parents: 82508
diff changeset
   389
          | assemble_app ((v, _) `|=> (t, _)) =
haftmann
parents: 82508
diff changeset
   390
              nbe_apps (nbe_abss 1 (ml_abs (ml_list [nbe_bound_optional v]) (assemble_iterm is_pat NONE t)))
haftmann
parents: 82508
diff changeset
   391
          | assemble_app (ICase { term = t, clauses = clauses, primitive = t0, ... }) =
haftmann
parents: 82508
diff changeset
   392
              nbe_apps (ml_cases (assemble_iterm is_pat NONE t)
haftmann
parents: 82508
diff changeset
   393
                (map (fn (p, t) => (assemble_iterm true NONE p, assemble_iterm is_pat a_match_fallback t)) clauses
haftmann
parents: 82508
diff changeset
   394
                  @ [("_", case a_match_fallback of SOME s => s | NONE => assemble_iterm is_pat NONE t0)]))
haftmann
parents: 82508
diff changeset
   395
        val (t', ts) = Code_Thingol.unfold_app t;
haftmann
parents: 82508
diff changeset
   396
        val a_ts = fold_rev (cons o assemble_iterm is_pat NONE) ts [];
haftmann
parents: 82508
diff changeset
   397
      in assemble_app t' a_ts end;
82375
haftmann
parents: 82374
diff changeset
   398
82508
haftmann
parents: 82507
diff changeset
   399
    fun assemble_fallback_fundef sym a_global_params ((samepairs, args), rhs) a_fallback_rhs =
28350
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28337
diff changeset
   400
      let
82509
haftmann
parents: 82508
diff changeset
   401
        val a_rhs_core = assemble_iterm false (SOME a_fallback_rhs) rhs;
82508
haftmann
parents: 82507
diff changeset
   402
        val a_rhs = if null samepairs then a_rhs_core
haftmann
parents: 82507
diff changeset
   403
          else ml_if (ml_and (map nbe_same samepairs)) a_rhs_core a_fallback_rhs;
haftmann
parents: 82507
diff changeset
   404
        val a_fallback_eqn = if forall Code_Thingol.is_IVar args then NONE
haftmann
parents: 82507
diff changeset
   405
          else SOME (replicate (length a_global_params) "_", a_fallback_rhs);
82509
haftmann
parents: 82508
diff changeset
   406
      in (map (assemble_iterm true NONE) args, a_rhs) :: the_list a_fallback_eqn end;
82506
289b18955960 revamped generation of functions
haftmann
parents: 82454
diff changeset
   407
82508
haftmann
parents: 82507
diff changeset
   408
    fun assemble_fallback_fundefs sym a_tparams a_dict_params a_global_params [(([], []), rhs)] =
82509
haftmann
parents: 82508
diff changeset
   409
          assemble_iterm false NONE rhs
82508
haftmann
parents: 82507
diff changeset
   410
      | assemble_fallback_fundefs sym a_tparams a_dict_params a_global_params eqns =
haftmann
parents: 82507
diff changeset
   411
          let
haftmann
parents: 82507
diff changeset
   412
            val a_fallback_syms = map_range (fn i => fallback_fun_ident i sym) (length eqns);
haftmann
parents: 82507
diff changeset
   413
            val a_fallback_rhss =
haftmann
parents: 82507
diff changeset
   414
              map_range (fn i => assemble_app_fallback_fun (i + 1) sym a_global_params) (length eqns - 1)
82510
haftmann
parents: 82509
diff changeset
   415
              @ [assemble_app_const sym a_tparams (a_dict_params @ a_global_params)];
82508
haftmann
parents: 82507
diff changeset
   416
          in
haftmann
parents: 82507
diff changeset
   417
            ml_let (ml_fundefs (a_fallback_syms ~~
haftmann
parents: 82507
diff changeset
   418
              map2 (assemble_fallback_fundef sym a_global_params) eqns a_fallback_rhss))
haftmann
parents: 82507
diff changeset
   419
              (assemble_app_fallback_fun 0 sym a_global_params)
haftmann
parents: 82507
diff changeset
   420
          end;
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   421
82508
haftmann
parents: 82507
diff changeset
   422
    fun assemble_fundef (sym, (num_args, (a_tparams, a_dict_params, a_global_params, eqns))) =
82506
289b18955960 revamped generation of functions
haftmann
parents: 82454
diff changeset
   423
      let
289b18955960 revamped generation of functions
haftmann
parents: 82454
diff changeset
   424
        val a_lhs = [ml_list a_tparams, ml_list (rev (a_dict_params @ a_global_params))];
82508
haftmann
parents: 82507
diff changeset
   425
        val a_rhs = assemble_fallback_fundefs sym a_tparams a_dict_params a_global_params eqns;
82506
289b18955960 revamped generation of functions
haftmann
parents: 82454
diff changeset
   426
        val a_univ = ml_abs (ml_list a_tparams) (nbe_abss num_args (assemble_fun sym a_tparams));
289b18955960 revamped generation of functions
haftmann
parents: 82454
diff changeset
   427
      in
289b18955960 revamped generation of functions
haftmann
parents: 82454
diff changeset
   428
        ((fun_ident sym, [(a_lhs, a_rhs), (a_lhs, ml_exc (fun_ident sym))]), a_univ)
289b18955960 revamped generation of functions
haftmann
parents: 82454
diff changeset
   429
      end;
82375
haftmann
parents: 82374
diff changeset
   430
82508
haftmann
parents: 82507
diff changeset
   431
    val (a_fun_defs, a_fun_vals) = map_split assemble_fundef eqnss;
82506
289b18955960 revamped generation of functions
haftmann
parents: 82454
diff changeset
   432
    val dep_params = ml_list (map fun_ident deps);
289b18955960 revamped generation of functions
haftmann
parents: 82454
diff changeset
   433
  in ml_abs dep_params (ml_let (ml_fundefs a_fun_defs) (ml_list a_fun_vals)) end;
25944
haftmann
parents: 25935
diff changeset
   434
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   435
fun assemble_eqnss ctxt idx_of_sym deps eqnss =
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   436
  assemble_preprocessed_eqnss ctxt idx_of_sym deps (map preprocess_eqns eqnss);
82375
haftmann
parents: 82374
diff changeset
   437
35371
6c92eb394e3c tuned whitespace
haftmann
parents: 34251
diff changeset
   438
63162
haftmann
parents: 63158
diff changeset
   439
(* compilation of equations *)
25944
haftmann
parents: 25935
diff changeset
   440
55757
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55167
diff changeset
   441
fun compile_eqnss ctxt nbe_program raw_deps [] = []
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55167
diff changeset
   442
  | compile_eqnss ctxt nbe_program raw_deps eqnss =
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   443
      let
26064
65585de05a66 using integers for pattern matching
haftmann
parents: 26011
diff changeset
   444
        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
   445
          (fn dep => Option.map (fn univ => (dep, univ)) (fst ((Code_Symbol.Graph.get_node nbe_program dep)))) raw_deps);
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   446
        val idx_of_sym = raw_deps
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   447
          |> 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
   448
          |> AList.lookup (op =)
65585de05a66 using integers for pattern matching
haftmann
parents: 26011
diff changeset
   449
          |> (fn f => the o f);
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   450
        val s = assemble_eqnss ctxt idx_of_sym deps eqnss;
82374
haftmann
parents: 82373
diff changeset
   451
        val syms = map fst eqnss;
25944
haftmann
parents: 25935
diff changeset
   452
      in
haftmann
parents: 25935
diff changeset
   453
        s
57435
312660c1a70a modernized diagnostic options
haftmann
parents: 57174
diff changeset
   454
        |> traced ctxt (fn s => "\n--- code to be evaluated:\n" ^ s)
38931
5e84c11c4b8a evaluate takes ml context and ml expression parameter
haftmann
parents: 38676
diff changeset
   455
        |> pair ""
39911
2b4430847310 moved ML_Context.value to Code_Runtime
haftmann
parents: 39606
diff changeset
   456
        |> Code_Runtime.value ctxt univs_cookie
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   457
        |> (fn dependent_fs => dependent_fs deps_vals)
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   458
        |> (fn fs => syms ~~ fs)
25944
haftmann
parents: 25935
diff changeset
   459
      end;
25190
5cd8486c5a4f clarified implementation
haftmann
parents: 25167
diff changeset
   460
30672
beaadd5af500 more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents: 30288
diff changeset
   461
63162
haftmann
parents: 63158
diff changeset
   462
(* extraction of equations from statements *)
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   463
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   464
fun dummy_const sym typargs dictss =
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   465
  IConst { sym = sym, typargs = typargs, dictss = dictss,
77233
6bdd125d932b explicit range types in abstractions
stuebinm <stuebinm@disroot.org>
parents: 74561
diff changeset
   466
    dom = [], annotation = NONE, range = ITyVar "" };
48072
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 47609
diff changeset
   467
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   468
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
   469
      []
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   470
  | eqns_of_stmt (_, Code_Thingol.Fun ((_, []), _)) =
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   471
      []
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   472
  | 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
   473
      [(sym_const, (vs, map fst eqns))]
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents: 27609
diff changeset
   474
  | eqns_of_stmt (_, Code_Thingol.Datatypecons _) =
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   475
      []
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents: 27609
diff changeset
   476
  | eqns_of_stmt (_, Code_Thingol.Datatype _) =
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   477
      []
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   478
  | eqns_of_stmt (sym_class, Code_Thingol.Class (v, (classrels, classparams))) =
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   479
      let
82444
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   480
        val syms = map (rpair [] o Class_Relation) classrels @ map (rpair [(v, [])] o Constant o fst) classparams;
81507
08574da77b4a clarified signature: shorten common cases;
wenzelm
parents: 80910
diff changeset
   481
        val params = Name.invent_global "d" (length syms);
82444
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   482
        fun mk (k, (sym, vs)) =
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   483
          (sym, (vs,
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   484
            [([dummy_const sym_class [] [] `$$ map (IVar o SOME) params],
31889
fb2c8a687529 all variable names are optional
haftmann
parents: 31888
diff changeset
   485
              IVar (SOME (nth params k)))]));
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   486
      in map_index mk syms end
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents: 27609
diff changeset
   487
  | eqns_of_stmt (_, Code_Thingol.Classrel _) =
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   488
      []
28054
2b84d34c5d02 restructured and split code serializer module
haftmann
parents: 27609
diff changeset
   489
  | eqns_of_stmt (_, Code_Thingol.Classparam _) =
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   490
      []
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   491
  | eqns_of_stmt (sym_inst, Code_Thingol.Classinst { class, tyco, vs, superinsts, inst_params, ... }) =
82444
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   492
      [(sym_inst, (vs, [([], dummy_const (Type_Class class) [] [] `$$
82447
741f6f6df144 reflect nested lists in variables names
haftmann
parents: 82444
diff changeset
   493
        map (fn (class, dictss) =>
741f6f6df144 reflect nested lists in variables names
haftmann
parents: 82444
diff changeset
   494
          dummy_const (Class_Instance (tyco, class)) (map (ITyVar o fst) vs) (map snd dictss)) superinsts
82444
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   495
          @ map (IConst o fst o snd o fst) inst_params)]))];
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   496
39392
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   497
63162
haftmann
parents: 63158
diff changeset
   498
(* compilation of whole programs *)
39392
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   499
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   500
fun ensure_sym_idx sym (nbe_program, (max_idx, sym_tab)) =
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   501
  if can (Code_Symbol.Graph.get_node nbe_program) sym
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   502
  then (nbe_program, (max_idx, sym_tab))
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   503
  else (Code_Symbol.Graph.new_node (sym, (NONE, max_idx)) nbe_program,
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   504
    (max_idx + 1, Inttab.update_new (max_idx, sym) sym_tab));
39399
267235a14938 static nbe conversion
haftmann
parents: 39396
diff changeset
   505
55757
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55167
diff changeset
   506
fun compile_stmts ctxt stmts_deps =
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   507
  let
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   508
    val names = map (fst o fst) stmts_deps;
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   509
    val names_deps = map (fn ((name, _), deps) => (name, deps)) stmts_deps;
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   510
    val eqnss = maps (eqns_of_stmt o fst) stmts_deps;
26064
65585de05a66 using integers for pattern matching
haftmann
parents: 26011
diff changeset
   511
    val refl_deps = names_deps
25190
5cd8486c5a4f clarified implementation
haftmann
parents: 25167
diff changeset
   512
      |> maps snd
5cd8486c5a4f clarified implementation
haftmann
parents: 25167
diff changeset
   513
      |> distinct (op =)
26064
65585de05a66 using integers for pattern matching
haftmann
parents: 26011
diff changeset
   514
      |> fold (insert (op =)) names;
39392
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   515
    fun compile nbe_program = eqnss
55757
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55167
diff changeset
   516
      |> compile_eqnss ctxt nbe_program refl_deps
39392
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   517
      |> rpair nbe_program;
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   518
  in
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   519
    fold ensure_sym_idx refl_deps
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   520
    #> 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
   521
      #> compile
82374
haftmann
parents: 82373
diff changeset
   522
      #-> fold (fn (sym, univ) => (Code_Symbol.Graph.map_node sym o apfst) (K (SOME univ))))
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   523
  end;
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   524
63164
72aaf69328fc optional timing for code generator conversions
haftmann
parents: 63162
diff changeset
   525
fun compile_program { ctxt, program } =
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   526
  let
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   527
    fun add_stmts names (nbe_program, (max_idx, sym_tab)) =
82375
haftmann
parents: 82374
diff changeset
   528
      if exists ((can o Code_Symbol.Graph.get_node) nbe_program) names
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   529
      then (nbe_program, (max_idx, sym_tab))
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   530
      else (nbe_program, (max_idx, sym_tab))
82374
haftmann
parents: 82373
diff changeset
   531
        |> compile_stmts ctxt (map (fn sym => ((sym, Code_Symbol.Graph.get_node program sym),
haftmann
parents: 82373
diff changeset
   532
          Code_Symbol.Graph.immediate_succs program sym)) names);
28706
3fef773ae6b1 restored incremental code generation
haftmann
parents: 28663
diff changeset
   533
  in
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   534
    fold_rev add_stmts (Code_Symbol.Graph.strong_conn program)
28706
3fef773ae6b1 restored incremental code generation
haftmann
parents: 28663
diff changeset
   535
  end;
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   536
25944
haftmann
parents: 25935
diff changeset
   537
63157
65a81a4ef7f8 clarified naming conventions and code for code evaluation sandwiches
haftmann
parents: 63156
diff changeset
   538
(** normalization **)
25944
haftmann
parents: 25935
diff changeset
   539
63162
haftmann
parents: 63158
diff changeset
   540
(* compilation and reconstruction of terms *)
25944
haftmann
parents: 25935
diff changeset
   541
82444
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   542
fun ad_hoc_eqn_of_term ((vs, _) : typscheme, t) =
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   543
  (Code_Symbol.value, (vs, [([], t)]));
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   544
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   545
fun compile_term { ctxt, nbe_program, deps, tfrees, vs_ty_t = vs_ty_t as ((vs, _), _) } =
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   546
  let
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   547
    val tparams = map (fn (v, _) => TParam v) tfrees;
25924
f974a1c64348 improved implementation
haftmann
parents: 25865
diff changeset
   548
    val dict_frees = maps (fn (v, sort) => map_index (curry DFree v o fst) sort) vs;
f974a1c64348 improved implementation
haftmann
parents: 25865
diff changeset
   549
  in
82444
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   550
    ad_hoc_eqn_of_term vs_ty_t
55757
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55167
diff changeset
   551
    |> singleton (compile_eqnss ctxt nbe_program deps)
25924
f974a1c64348 improved implementation
haftmann
parents: 25865
diff changeset
   552
    |> snd
82444
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   553
    |> (fn f => apps (f tparams) (rev dict_frees))
25924
f974a1c64348 improved implementation
haftmann
parents: 25865
diff changeset
   554
  end;
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   555
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   556
fun reconstruct_term ctxt sym_tab tfrees =
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   557
  let
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   558
    fun take_until f [] = []
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   559
      | take_until f (x :: xs) = if f x then [] else x :: take_until f xs;
82444
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   560
    fun is_dict (Const ((idx, _), _)) =
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   561
          (case Inttab.lookup sym_tab idx of
55150
0940309ed8f1 less clumsy namespace
haftmann
parents: 55147
diff changeset
   562
            SOME (Constant _) => false
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   563
          | _ => true)
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   564
      | is_dict (DFree _) = true
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   565
      | is_dict _ = false;
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   566
    fun const_of_idx idx =
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   567
      case Inttab.lookup sym_tab idx of SOME (Constant const) => const;
82444
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   568
    fun reconstruct_type (Type (tyco, tys)) = Term.Type (tyco, map reconstruct_type tys)
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   569
      | reconstruct_type (TParam v) = TFree (v, the (AList.lookup (op =) tfrees v));
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   570
    fun of_apps bounds (t, ts) =
82444
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   571
      list_comb (t, rev (map (of_univ bounds) ts))
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   572
    and of_univ bounds (Const ((idx, tparams), us)) =
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   573
          let
82375
haftmann
parents: 82374
diff changeset
   574
            val const = const_of_idx idx;
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   575
            val us' = take_until is_dict us;
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   576
            val T = Consts.instance (Proof_Context.consts_of ctxt) (const, map reconstruct_type tparams);
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   577
          in of_apps bounds (Term.Const (const, T), us') end
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   578
      | of_univ bounds (BVar (i, us)) =
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   579
          of_apps bounds (Bound (bounds - i - 1), us)
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   580
      | of_univ bounds (u as Abs _) =
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   581
          Term.Abs ("u", dummyT, of_univ (bounds + 1) (apps u [BVar (bounds, [])]))
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   582
  in of_univ 0 end;
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   583
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   584
fun compile_and_reconstruct_term { ctxt, nbe_program, sym_tab, deps, tfrees, vs_ty_t } =
82444
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   585
  compile_term { ctxt = ctxt, nbe_program = nbe_program, deps = deps, tfrees = tfrees, vs_ty_t = vs_ty_t }
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   586
  |> reconstruct_term ctxt sym_tab tfrees;
35371
6c92eb394e3c tuned whitespace
haftmann
parents: 34251
diff changeset
   587
82443
3e92066d2be7 restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents: 82442
diff changeset
   588
fun retype_term ctxt t T =
3e92066d2be7 restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents: 82442
diff changeset
   589
  let
3e92066d2be7 restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents: 82442
diff changeset
   590
    val ctxt' =
3e92066d2be7 restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents: 82442
diff changeset
   591
      ctxt
3e92066d2be7 restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents: 82442
diff changeset
   592
      |> Variable.declare_typ T
3e92066d2be7 restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents: 82442
diff changeset
   593
      |> Config.put Type_Infer.object_logic false
3e92066d2be7 restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents: 82442
diff changeset
   594
      |> Config.put Type_Infer_Context.const_sorts false
3e92066d2be7 restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents: 82442
diff changeset
   595
  in
3e92066d2be7 restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents: 82442
diff changeset
   596
    singleton (Variable.export_terms ctxt' ctxt') (Syntax.check_term ctxt' (Type.constraint T t))
3e92066d2be7 restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents: 82442
diff changeset
   597
  end;
3e92066d2be7 restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents: 82442
diff changeset
   598
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   599
fun normalize_term (nbe_program, sym_tab) raw_ctxt t_original vs_ty_t deps =
39392
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   600
  let
82444
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   601
    val T = fastype_of t_original;
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   602
    val tfrees = Term.add_tfrees t_original [];
55757
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55167
diff changeset
   603
    val ctxt = Syntax.init_pretty_global (Proof_Context.theory_of raw_ctxt);
82443
3e92066d2be7 restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents: 82442
diff changeset
   604
    val string_of_term =
3e92066d2be7 restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents: 82442
diff changeset
   605
      Syntax.string_of_term
3e92066d2be7 restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents: 82442
diff changeset
   606
         (ctxt
3e92066d2be7 restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents: 82442
diff changeset
   607
           |> Config.put show_types true
3e92066d2be7 restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents: 82442
diff changeset
   608
           |> Config.put show_sorts true);
82444
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   609
    fun retype t' = retype_term ctxt t' T;
56969
7491932da574 dropped obsolete hand-waving adjustion of type variables: safely done in preprocessor
haftmann
parents: 56925
diff changeset
   610
    fun check_tvars t' =
7491932da574 dropped obsolete hand-waving adjustion of type variables: safely done in preprocessor
haftmann
parents: 56925
diff changeset
   611
      if null (Term.add_tvars t' []) then t'
7491932da574 dropped obsolete hand-waving adjustion of type variables: safely done in preprocessor
haftmann
parents: 56925
diff changeset
   612
      else error ("Illegal schematic type variables in normalized term: " ^ string_of_term t');
39392
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   613
  in
63164
72aaf69328fc optional timing for code generator conversions
haftmann
parents: 63162
diff changeset
   614
    Code_Preproc.timed "computing NBE expression" #ctxt compile_and_reconstruct_term
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   615
      { ctxt = ctxt, nbe_program = nbe_program, sym_tab = sym_tab, deps = deps,
82444
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   616
        tfrees = tfrees, vs_ty_t = vs_ty_t }
57435
312660c1a70a modernized diagnostic options
haftmann
parents: 57174
diff changeset
   617
    |> traced ctxt (fn t => "Normalized:\n" ^ string_of_term t)
82443
3e92066d2be7 restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents: 82442
diff changeset
   618
    |> retype
57435
312660c1a70a modernized diagnostic options
haftmann
parents: 57174
diff changeset
   619
    |> traced ctxt (fn t => "Types inferred:\n" ^ string_of_term t)
39392
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   620
    |> check_tvars
57435
312660c1a70a modernized diagnostic options
haftmann
parents: 57174
diff changeset
   621
    |> traced ctxt (fn _ => "---\n")
39392
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   622
  end;
7a0fcee7a2a3 more clear separation of static compilation and dynamic evaluation
haftmann
parents: 39388
diff changeset
   623
39436
4a7d09da2b9c tuned whitespace
haftmann
parents: 39399
diff changeset
   624
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   625
(* function store *)
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   626
34173
458ced35abb8 reduced code generator cache to the baremost minimum
haftmann
parents: 33522
diff changeset
   627
structure Nbe_Functions = Code_Data
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   628
(
82444
7a9164068583 represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents: 82443
diff changeset
   629
  type T = ((Type list -> Univ) option * int) Code_Symbol.Graph.T * (int * Code_Symbol.T Inttab.table);
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 55043
diff changeset
   630
  val empty = (Code_Symbol.Graph.empty, (0, Inttab.empty));
25101
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   631
);
cae0f68b693b now employing dictionaries
haftmann
parents: 25098
diff changeset
   632
55757
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55167
diff changeset
   633
fun compile ignore_cache ctxt program =
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   634
  let
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   635
    val (nbe_program, (_, sym_tab)) =
55757
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55167
diff changeset
   636
      Nbe_Functions.change (if ignore_cache then NONE else SOME (Proof_Context.theory_of ctxt))
63164
72aaf69328fc optional timing for code generator conversions
haftmann
parents: 63162
diff changeset
   637
        (Code_Preproc.timed "compiling NBE program" #ctxt
72aaf69328fc optional timing for code generator conversions
haftmann
parents: 63162
diff changeset
   638
          compile_program { ctxt = ctxt, program = program });
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   639
  in (nbe_program, sym_tab) end;
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   640
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32123
diff changeset
   641
47572
1e18bbfb40cb tuned heading
haftmann
parents: 44789
diff changeset
   642
(* evaluation oracle *)
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   643
55757
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55167
diff changeset
   644
fun mk_equals ctxt lhs raw_rhs =
26747
f32fa5f5bdd1 moved 'trivial classes' to foundation of code generator
haftmann
parents: 26739
diff changeset
   645
  let
59586
ddf6deaadfe8 clarified signature;
wenzelm
parents: 59582
diff changeset
   646
    val ty = Thm.typ_of_cterm lhs;
74295
9a9326a072bb more antiquotations;
wenzelm
parents: 74282
diff changeset
   647
    val eq = Thm.cterm_of ctxt \<^Const>\<open>Pure.eq ty\<close>;
59642
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   648
    val rhs = Thm.cterm_of ctxt raw_rhs;
30947
dd551284a300 re-engineering of evaluation conversions
haftmann
parents: 30942
diff changeset
   649
  in Thm.mk_binop eq lhs rhs end;
26747
f32fa5f5bdd1 moved 'trivial classes' to foundation of code generator
haftmann
parents: 26739
diff changeset
   650
78795
f7e972d567f3 clarified signature: more concise variations on implicit theory setup;
wenzelm
parents: 77707
diff changeset
   651
val (_, raw_oracle) =
f7e972d567f3 clarified signature: more concise variations on implicit theory setup;
wenzelm
parents: 77707
diff changeset
   652
  Theory.setup_result (Thm.add_oracle (\<^binding>\<open>normalization_by_evaluation\<close>,
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   653
    fn (nbe_program_sym_tab, ctxt, vs_ty_t, deps, ct) =>
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   654
      mk_equals ctxt ct (normalize_term nbe_program_sym_tab ctxt (Thm.term_of ct) vs_ty_t deps)));
55757
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55167
diff changeset
   655
82454
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   656
fun oracle nbe_program_sym_tab ctxt vs_ty_t deps ct =
ba1f9fb23b8d clarified variable names
haftmann
parents: 82453
diff changeset
   657
  raw_oracle (nbe_program_sym_tab, ctxt, vs_ty_t, deps, ct);
31064
ce37d8f48a9f treat frees driectly by the LCF kernel
haftmann
parents: 31049
diff changeset
   658
55757
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55167
diff changeset
   659
fun dynamic_conv ctxt = lift_triv_classes_conv ctxt
63158
534f16b0ca39 explicit quasi-global context for nbe conversions -- works around quasi-global type variable handling in lift_triv_classes_conv
haftmann
parents: 63157
diff changeset
   660
  (fn ctxt' => Code_Thingol.dynamic_conv ctxt' (fn program =>
534f16b0ca39 explicit quasi-global context for nbe conversions -- works around quasi-global type variable handling in lift_triv_classes_conv
haftmann
parents: 63157
diff changeset
   661
    oracle (compile false ctxt program) ctxt'));
30942
1e246776f876 diagnostic commands now in code_thingol; tuned code of funny continuations
haftmann
parents: 30672
diff changeset
   662
55757
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55167
diff changeset
   663
fun dynamic_value ctxt = lift_triv_classes_rew ctxt
63157
65a81a4ef7f8 clarified naming conventions and code for code evaluation sandwiches
haftmann
parents: 63156
diff changeset
   664
  (Code_Thingol.dynamic_value ctxt I (fn program =>
63162
haftmann
parents: 63158
diff changeset
   665
    normalize_term (compile false ctxt program) ctxt));
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   666
56973
62da80041afd syntactic means to prevent accidental mixup of static and dynamic context
haftmann
parents: 56972
diff changeset
   667
fun static_conv (ctxt_consts as { ctxt, ... }) =
55757
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55167
diff changeset
   668
  let
63157
65a81a4ef7f8 clarified naming conventions and code for code evaluation sandwiches
haftmann
parents: 63156
diff changeset
   669
    val conv = Code_Thingol.static_conv_thingol ctxt_consts
63164
72aaf69328fc optional timing for code generator conversions
haftmann
parents: 63162
diff changeset
   670
      (fn { program, deps = _ } => oracle (compile true ctxt program));
63158
534f16b0ca39 explicit quasi-global context for nbe conversions -- works around quasi-global type variable handling in lift_triv_classes_conv
haftmann
parents: 63157
diff changeset
   671
  in fn ctxt' => lift_triv_classes_conv ctxt' conv end;
39399
267235a14938 static nbe conversion
haftmann
parents: 39396
diff changeset
   672
56973
62da80041afd syntactic means to prevent accidental mixup of static and dynamic context
haftmann
parents: 56972
diff changeset
   673
fun static_value { ctxt, consts } =
55757
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55167
diff changeset
   674
  let
63157
65a81a4ef7f8 clarified naming conventions and code for code evaluation sandwiches
haftmann
parents: 63156
diff changeset
   675
    val comp = Code_Thingol.static_value { ctxt = ctxt, lift_postproc = I, consts = consts }
63164
72aaf69328fc optional timing for code generator conversions
haftmann
parents: 63162
diff changeset
   676
      (fn { program, deps = _ } => normalize_term (compile false ctxt program));
63157
65a81a4ef7f8 clarified naming conventions and code for code evaluation sandwiches
haftmann
parents: 63156
diff changeset
   677
  in fn ctxt' => lift_triv_classes_rew ctxt' (comp ctxt') end;
41184
5c6f44d22f51 simplified evaluation function names
haftmann
parents: 41118
diff changeset
   678
24155
d86867645f4f nbe improved
haftmann
parents:
diff changeset
   679
end;