src/Pure/Tools/codegen_funcgr.ML
author haftmann
Thu, 04 Jan 2007 14:01:39 +0100
changeset 21989 0315ecfd3d5d
parent 21915 4e63c55f4cb4
child 22021 6466a24dee5b
permissions -rw-r--r--
eta-expansion now only to common maximum number of arguments
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
     1
(*  Title:      Pure/Tools/codegen_funcgr.ML
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
     2
    ID:         $Id$
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
     3
    Author:     Florian Haftmann, TU Muenchen
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
     4
20855
9f60d493c8fe clarified header comments
haftmann
parents: 20835
diff changeset
     5
Retrieving, normalizing and structuring code function theorems
9f60d493c8fe clarified header comments
haftmann
parents: 20835
diff changeset
     6
in graph with explicit dependencies.
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
     7
*)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
     8
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
     9
signature CODEGEN_FUNCGR =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    10
sig
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    11
  type T;
21120
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
    12
  val make: theory -> CodegenConsts.const list -> T
21129
8e621232a865 clarified make_term interface
haftmann
parents: 21120
diff changeset
    13
  val make_term: theory -> ((thm -> thm) -> cterm -> thm -> 'a) -> cterm -> 'a * T
21120
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
    14
  val funcs: T -> CodegenConsts.const -> thm list
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
    15
  val typ: T -> CodegenConsts.const -> typ
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
    16
  val deps: T -> CodegenConsts.const list -> CodegenConsts.const list list
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
    17
  val all: T -> CodegenConsts.const list
21989
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    18
  val norm_varnames: theory -> thm list -> thm list
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    19
  val expand_eta: theory -> int -> thm -> thm
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    20
  val print_codethms: theory -> CodegenConsts.const list -> unit
20938
041badc7fcaf added keeping of funcgr
haftmann
parents: 20896
diff changeset
    21
  structure Constgraph : GRAPH
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    22
end;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    23
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    24
structure CodegenFuncgr: CODEGEN_FUNCGR =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    25
struct
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    26
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    27
(** code data **)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    28
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    29
structure Consttab = CodegenConsts.Consttab;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    30
structure Constgraph = GraphFun (
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    31
  type key = CodegenConsts.const;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    32
  val ord = CodegenConsts.const_ord;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    33
);
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    34
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    35
type T = (typ * thm list) Constgraph.T;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    36
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    37
structure Funcgr = CodeDataFun
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    38
(struct
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    39
  val name = "Pure/codegen_funcgr";
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    40
  type T = T;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    41
  val empty = Constgraph.empty;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    42
  fun merge _ _ = Constgraph.empty;
20938
041badc7fcaf added keeping of funcgr
haftmann
parents: 20896
diff changeset
    43
  fun purge _ NONE _ = Constgraph.empty
041badc7fcaf added keeping of funcgr
haftmann
parents: 20896
diff changeset
    44
    | purge _ (SOME cs) funcgr =
21463
42dd50268c8b completed class parameter handling in axclass.ML
haftmann
parents: 21387
diff changeset
    45
        Constgraph.del_nodes ((Constgraph.all_preds funcgr 
20938
041badc7fcaf added keeping of funcgr
haftmann
parents: 20896
diff changeset
    46
          o filter (can (Constgraph.get_node funcgr))) cs) funcgr;
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    47
end);
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    48
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    49
val _ = Context.add_setup Funcgr.init;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    50
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    51
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    52
(** theorem purification **)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    53
21989
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    54
fun expand_eta thy k thm =
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    55
  let
21989
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    56
    val (lhs, rhs) = (Logic.dest_equals o Drule.plain_prop_of) thm;
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    57
    val (head, args) = strip_comb lhs;
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    58
    val l = Int.max (0, k - length args);
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    59
    val used = Name.make_context (map (fst o fst) (Term.add_vars lhs []));
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    60
    fun get_name _ 0 used = ([], used)
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    61
      | get_name (Abs (v, ty, t)) k used =
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    62
          used
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    63
          |> Name.variants [CodegenNames.purify_var v]
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    64
          ||>> get_name t (k - 1)
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    65
          |>> (fn ([v'], vs') => (v', ty) :: vs')
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    66
      | get_name t k used = 
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    67
          let
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    68
            val (tys, _) = (strip_type o fastype_of) t
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    69
          in case tys
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    70
           of [] => raise TERM ("expand_eta", [t])
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    71
            | ty :: _ =>
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    72
                used
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    73
                |> Name.variants [""]
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    74
                |-> (fn [v] => get_name (t $ Var ((v, 0), ty)) (k - 1)
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    75
                #>> (fn vs' => (v, ty) :: vs'))
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    76
          end;
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    77
    val (vs, _) = get_name rhs l used;
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    78
    val vs_refl = map (fn (v, ty) => Thm.reflexive (Thm.cterm_of thy (Var ((v, 0), ty)))) vs;
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    79
  in
21989
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    80
    fold (fn refl => fn thm => Thm.combination thm refl) vs_refl thm
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    81
  end;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    82
21989
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    83
fun norm_args thy thms =
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    84
  let
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    85
    val num_args_of = length o snd o strip_comb o fst o Logic.dest_equals;
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    86
    val k = fold (curry Int.max o num_args_of o Drule.plain_prop_of) thms 0;
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    87
  in
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    88
    thms
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    89
    |> map (expand_eta thy k)
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    90
    |> map (Drule.fconv_rule Drule.beta_eta_conversion)
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
    91
  end;
21120
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
    92
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    93
fun canonical_tvars thy thm =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    94
  let
21159
7f6bdffe3d06 fixed problem with variable names
haftmann
parents: 21129
diff changeset
    95
    fun tvars_subst_for thm = (fold_types o fold_atyps)
7f6bdffe3d06 fixed problem with variable names
haftmann
parents: 21129
diff changeset
    96
      (fn TVar (v_i as (v, _), sort) => let
21915
4e63c55f4cb4 different handling of type variable names
haftmann
parents: 21463
diff changeset
    97
            val v' = CodegenNames.purify_tvar v
21159
7f6bdffe3d06 fixed problem with variable names
haftmann
parents: 21129
diff changeset
    98
          in if v = v' then I
7f6bdffe3d06 fixed problem with variable names
haftmann
parents: 21129
diff changeset
    99
          else insert (op =) (v_i, (v', sort)) end
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   100
        | _ => I) (prop_of thm) [];
21159
7f6bdffe3d06 fixed problem with variable names
haftmann
parents: 21129
diff changeset
   101
    fun mk_inst (v_i, (v', sort)) (maxidx, acc) =
7f6bdffe3d06 fixed problem with variable names
haftmann
parents: 21129
diff changeset
   102
      let
7f6bdffe3d06 fixed problem with variable names
haftmann
parents: 21129
diff changeset
   103
        val ty = TVar (v_i, sort)
7f6bdffe3d06 fixed problem with variable names
haftmann
parents: 21129
diff changeset
   104
      in
7f6bdffe3d06 fixed problem with variable names
haftmann
parents: 21129
diff changeset
   105
        (maxidx + 1, (ctyp_of thy ty, ctyp_of thy (TVar ((v', maxidx), sort))) :: acc)
7f6bdffe3d06 fixed problem with variable names
haftmann
parents: 21129
diff changeset
   106
      end;
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   107
    val maxidx = Thm.maxidx_of thm + 1;
21159
7f6bdffe3d06 fixed problem with variable names
haftmann
parents: 21129
diff changeset
   108
    val (_, inst) = fold mk_inst (tvars_subst_for thm) (maxidx + 1, []);
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   109
  in Thm.instantiate (inst, []) thm end;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   110
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   111
fun canonical_vars thy thm =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   112
  let
21159
7f6bdffe3d06 fixed problem with variable names
haftmann
parents: 21129
diff changeset
   113
    fun vars_subst_for thm = fold_aterms
7f6bdffe3d06 fixed problem with variable names
haftmann
parents: 21129
diff changeset
   114
      (fn Var (v_i as (v, _), ty) => let
7f6bdffe3d06 fixed problem with variable names
haftmann
parents: 21129
diff changeset
   115
            val v' = CodegenNames.purify_var v
7f6bdffe3d06 fixed problem with variable names
haftmann
parents: 21129
diff changeset
   116
          in if v = v' then I
7f6bdffe3d06 fixed problem with variable names
haftmann
parents: 21129
diff changeset
   117
          else insert (op =) (v_i, (v', ty)) end
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   118
        | _ => I) (prop_of thm) [];
21159
7f6bdffe3d06 fixed problem with variable names
haftmann
parents: 21129
diff changeset
   119
    fun mk_inst (v_i as (v, i), (v', ty)) (maxidx, acc) =
7f6bdffe3d06 fixed problem with variable names
haftmann
parents: 21129
diff changeset
   120
      let
7f6bdffe3d06 fixed problem with variable names
haftmann
parents: 21129
diff changeset
   121
        val t = Var (v_i, ty)
7f6bdffe3d06 fixed problem with variable names
haftmann
parents: 21129
diff changeset
   122
      in
7f6bdffe3d06 fixed problem with variable names
haftmann
parents: 21129
diff changeset
   123
        (maxidx + 1, (cterm_of thy t, cterm_of thy (Var ((v', maxidx), ty))) :: acc)
7f6bdffe3d06 fixed problem with variable names
haftmann
parents: 21129
diff changeset
   124
      end;
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   125
    val maxidx = Thm.maxidx_of thm + 1;
21159
7f6bdffe3d06 fixed problem with variable names
haftmann
parents: 21129
diff changeset
   126
    val (_, inst) = fold mk_inst (vars_subst_for thm) (maxidx + 1, []);
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   127
  in Thm.instantiate ([], inst) thm end;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   128
21989
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
   129
fun norm_varnames thy thms =
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   130
  let
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   131
    fun burrow_thms f [] = []
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   132
      | burrow_thms f thms =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   133
          thms
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   134
          |> Conjunction.intr_list
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   135
          |> f
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   136
          |> Conjunction.elim_list;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   137
  in
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   138
    thms
21989
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
   139
    |> norm_args thy
21159
7f6bdffe3d06 fixed problem with variable names
haftmann
parents: 21129
diff changeset
   140
    |> burrow_thms (canonical_tvars thy)
7f6bdffe3d06 fixed problem with variable names
haftmann
parents: 21129
diff changeset
   141
    |> map (canonical_vars thy)
7f6bdffe3d06 fixed problem with variable names
haftmann
parents: 21129
diff changeset
   142
    |> map Drule.zero_var_indexes
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   143
  end;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   144
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   145
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   146
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   147
(** retrieval **)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   148
21120
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   149
fun funcs funcgr =
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   150
  these o Option.map snd o try (Constgraph.get_node funcgr);
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   151
21120
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   152
fun typ funcgr =
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   153
  fst o Constgraph.get_node funcgr;
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   154
21120
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   155
fun deps funcgr cs =
20835
haftmann
parents: 20705
diff changeset
   156
  let
haftmann
parents: 20705
diff changeset
   157
    val conn = Constgraph.strong_conn funcgr;
haftmann
parents: 20705
diff changeset
   158
    val order = rev conn;
haftmann
parents: 20705
diff changeset
   159
  in
haftmann
parents: 20705
diff changeset
   160
    (map o filter) (member (op =) (Constgraph.all_succs funcgr cs)) order
haftmann
parents: 20705
diff changeset
   161
    |> filter_out null
haftmann
parents: 20705
diff changeset
   162
  end;
haftmann
parents: 20705
diff changeset
   163
21120
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   164
fun all funcgr = Constgraph.keys funcgr;
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   165
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   166
local
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   167
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   168
fun add_things_of thy f (c, thms) =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   169
  (fold o fold_aterms)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   170
     (fn Const c_ty => let
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   171
            val c' = CodegenConsts.norm_of_typ thy c_ty
21120
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   172
          in if is_some c andalso CodegenConsts.eq_const (the c, c') then I
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   173
          else f (c', c_ty) end
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   174
       | _ => I) (maps (op :: o swap o apfst (snd o strip_comb)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   175
            o Logic.dest_equals o Drule.plain_prop_of) thms)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   176
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   177
fun rhs_of thy (c, thms) =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   178
  Consttab.empty
21120
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   179
  |> add_things_of thy (Consttab.update o rpair () o fst) (SOME c, thms)
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   180
  |> Consttab.keys;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   181
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   182
fun insts_of thy funcgr (c, ty) =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   183
  let
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   184
    val tys = Sign.const_typargs thy (c, ty);
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   185
    val c' = CodegenConsts.norm thy (c, tys);
20705
da71d46b8b2f fixed some mess
haftmann
parents: 20600
diff changeset
   186
    val ty_decl = CodegenConsts.disc_typ_of_const thy
da71d46b8b2f fixed some mess
haftmann
parents: 20600
diff changeset
   187
      (fst o Constgraph.get_node funcgr o CodegenConsts.norm thy) (c, tys);
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   188
    val tys_decl = Sign.const_typargs thy (c, ty_decl);
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   189
    fun constructor tyco xs class =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   190
      (tyco, class) :: maps (maps fst) xs;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   191
    fun variable (TVar (_, sort)) = map (pair []) sort
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   192
      | variable (TFree (_, sort)) = map (pair []) sort;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   193
    fun mk_inst ty (TVar (_, sort)) = cons (ty, sort)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   194
      | mk_inst ty (TFree (_, sort)) = cons (ty, sort)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   195
      | mk_inst (Type (tyco1, tys1)) (Type (tyco2, tys2)) =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   196
          if tyco1 <> tyco2 then error "bad instance"
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   197
          else fold2 mk_inst tys1 tys2;
21120
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   198
    val pp = Sign.pp thy;
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   199
    val algebra = Sign.classes_of thy;
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   200
    fun classrel (x, _) _ = x;
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   201
    fun of_sort_deriv (ty, sort) =
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   202
      Sorts.of_sort_derivation pp algebra
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   203
        { classrel = classrel, constructor = constructor, variable = variable }
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   204
        (ty, sort)
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   205
  in
21120
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   206
    flat (maps of_sort_deriv (fold2 mk_inst tys tys_decl []))
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   207
  end;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   208
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   209
fun all_classops thy tyco class =
21463
42dd50268c8b completed class parameter handling in axclass.ML
haftmann
parents: 21387
diff changeset
   210
  try (AxClass.params_of_class thy) class
42dd50268c8b completed class parameter handling in axclass.ML
haftmann
parents: 21387
diff changeset
   211
  |> Option.map snd
42dd50268c8b completed class parameter handling in axclass.ML
haftmann
parents: 21387
diff changeset
   212
  |> these
42dd50268c8b completed class parameter handling in axclass.ML
haftmann
parents: 21387
diff changeset
   213
  |> map (fn (c, _) => (c, CodegenConsts.disc_typ_of_classop thy (c, [Type (tyco, [])])))
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   214
  |> map (CodegenConsts.norm_of_typ thy);
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   215
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   216
fun instdefs_of thy insts =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   217
  let
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   218
    val thy_classes = (#classes o Sorts.rep_algebra o Sign.classes_of) thy;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   219
  in
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   220
    Symtab.empty
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   221
    |> fold (fn (tyco, class) =>
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   222
        Symtab.map_default (tyco, []) (insert (op =) class)) insts
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   223
    |> (fn tab => Symtab.fold (fn (tyco, classes) => append (maps (all_classops thy tyco)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   224
         (Graph.all_succs thy_classes classes))) tab [])
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   225
  end;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   226
21120
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   227
fun insts_of_thms thy funcgr (c, thms) =
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   228
  let
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   229
    val insts = add_things_of thy (fn (_, c_ty) => fold (insert (op =))
21120
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   230
      (insts_of thy funcgr c_ty)) (SOME c, thms) [];
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   231
  in instdefs_of thy insts end;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   232
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   233
fun ensure_const thy funcgr c auxgr =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   234
  if can (Constgraph.get_node funcgr) c
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   235
    then (NONE, auxgr)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   236
  else if can (Constgraph.get_node auxgr) c
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   237
    then (SOME c, auxgr)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   238
  else if is_some (CodegenData.get_datatype_of_constr thy c) then
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   239
    auxgr
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   240
    |> Constgraph.new_node (c, [])
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   241
    |> pair (SOME c)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   242
  else let
21989
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
   243
    val thms = norm_varnames thy (CodegenData.these_funcs thy c);
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   244
    val rhs = rhs_of thy (c, thms);
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   245
  in
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   246
    auxgr
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   247
    |> Constgraph.new_node (c, thms)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   248
    |> fold_map (ensure_const thy funcgr) rhs
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   249
    |-> (fn rhs' => fold (fn SOME c' => Constgraph.add_edge (c, c')
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   250
                           | NONE => I) rhs')
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   251
    |> pair (SOME c)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   252
  end;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   253
21387
5d3d340cb783 clarified code for building function equation system; explicit check of type discipline
haftmann
parents: 21196
diff changeset
   254
exception INVALID of CodegenConsts.const list * string;
5d3d340cb783 clarified code for building function equation system; explicit check of type discipline
haftmann
parents: 21196
diff changeset
   255
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   256
fun specialize_typs thy funcgr eqss =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   257
  let
21387
5d3d340cb783 clarified code for building function equation system; explicit check of type discipline
haftmann
parents: 21196
diff changeset
   258
    fun max [] = 0
5d3d340cb783 clarified code for building function equation system; explicit check of type discipline
haftmann
parents: 21196
diff changeset
   259
      | max [l] = l
5d3d340cb783 clarified code for building function equation system; explicit check of type discipline
haftmann
parents: 21196
diff changeset
   260
      | max (k::l::ls) = max ((if k < l then l else k) :: ls);
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   261
    fun typscheme_of (c, ty) =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   262
      try (Constgraph.get_node funcgr) (CodegenConsts.norm_of_typ thy (c, ty))
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   263
      |> Option.map fst;
21387
5d3d340cb783 clarified code for building function equation system; explicit check of type discipline
haftmann
parents: 21196
diff changeset
   264
    fun incr_indices (c, thms) maxidx =
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   265
      let
21387
5d3d340cb783 clarified code for building function equation system; explicit check of type discipline
haftmann
parents: 21196
diff changeset
   266
        val thms' = map (Thm.incr_indexes (maxidx + 1)) thms;
5d3d340cb783 clarified code for building function equation system; explicit check of type discipline
haftmann
parents: 21196
diff changeset
   267
        val maxidx' = max (maxidx :: map Thm.maxidx_of thms');
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   268
      in ((c, thms'), maxidx') end;
21387
5d3d340cb783 clarified code for building function equation system; explicit check of type discipline
haftmann
parents: 21196
diff changeset
   269
    val (eqss', maxidx) =
5d3d340cb783 clarified code for building function equation system; explicit check of type discipline
haftmann
parents: 21196
diff changeset
   270
      fold_map incr_indices eqss 0;
5d3d340cb783 clarified code for building function equation system; explicit check of type discipline
haftmann
parents: 21196
diff changeset
   271
    fun unify_const (c, ty) (env, maxidx) =
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   272
      case typscheme_of (c, ty)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   273
       of SOME ty_decl => let
21387
5d3d340cb783 clarified code for building function equation system; explicit check of type discipline
haftmann
parents: 21196
diff changeset
   274
            val ty_decl' = Logic.incr_tvar (maxidx + 1) ty_decl;
5d3d340cb783 clarified code for building function equation system; explicit check of type discipline
haftmann
parents: 21196
diff changeset
   275
            val maxidx' = max [maxidx, Term.maxidx_of_typ ty_decl'];
5d3d340cb783 clarified code for building function equation system; explicit check of type discipline
haftmann
parents: 21196
diff changeset
   276
          in Type.unify (Sign.tsig_of thy) (ty_decl', ty) (env, maxidx')
5d3d340cb783 clarified code for building function equation system; explicit check of type discipline
haftmann
parents: 21196
diff changeset
   277
          handle TUNIFY => raise INVALID ([], setmp show_sorts true (setmp show_types true (fn f => f ())) (fn _ => ("Failed to instantiate\n"
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   278
            ^ (Sign.string_of_typ thy o Envir.norm_type env) ty_decl' ^ "\nto\n"
21387
5d3d340cb783 clarified code for building function equation system; explicit check of type discipline
haftmann
parents: 21196
diff changeset
   279
            ^ (Sign.string_of_typ thy o Envir.norm_type env) ty
5d3d340cb783 clarified code for building function equation system; explicit check of type discipline
haftmann
parents: 21196
diff changeset
   280
            ^ ",\nfor constant " ^ quote c
5d3d340cb783 clarified code for building function equation system; explicit check of type discipline
haftmann
parents: 21196
diff changeset
   281
            ^ "\nin function theorems\n"
5d3d340cb783 clarified code for building function equation system; explicit check of type discipline
haftmann
parents: 21196
diff changeset
   282
            ^ (cat_lines o maps (map (Sign.string_of_term thy o map_types (Envir.norm_type env) o Thm.prop_of) o snd)) eqss')))
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   283
          end
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   284
        | NONE => (env, maxidx);
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   285
    fun apply_unifier unif (c, []) = (c, [])
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   286
      | apply_unifier unif (c, thms as thm :: _) =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   287
          let
21120
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   288
            val tvars = Term.add_tvars (Thm.prop_of thm) [];
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   289
            fun mk_inst (v_i_sort as (v, _)) =
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   290
              let
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   291
                val ty = TVar v_i_sort;
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   292
              in
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   293
                pairself (Thm.ctyp_of thy) (ty,
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   294
                  TVar (v, (snd o dest_TVar o Envir.norm_type unif) ty))
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   295
              end;
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   296
            val instmap = map mk_inst tvars;
21387
5d3d340cb783 clarified code for building function equation system; explicit check of type discipline
haftmann
parents: 21196
diff changeset
   297
            val (thms' as thm' :: _) = map (Drule.zero_var_indexes o Thm.instantiate (instmap, [])) thms;
21989
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
   298
            val _ = if fst c = "" then ()
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
   299
              else case pairself (CodegenData.typ_func thy) (thm, thm')
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
   300
               of (ty, ty') => if (is_none o AxClass.class_of_param thy o fst) c
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
   301
                  andalso Sign.typ_equiv thy (Type.strip_sorts ty, Type.strip_sorts ty')
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
   302
                  orelse Sign.typ_equiv thy (ty, ty')
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
   303
                  then ()
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
   304
                  else raise INVALID ([], "illegal function type instantiation:\n" ^ Sign.string_of_typ thy (CodegenData.typ_func thy thm)
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
   305
                    ^ "\nto " ^ Sign.string_of_typ thy (CodegenData.typ_func thy thm')
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
   306
                    ^ ",\nfor constant " ^ CodegenConsts.string_of_const thy c
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
   307
                    ^ "\nin function theorems\n"
0315ecfd3d5d eta-expansion now only to common maximum number of arguments
haftmann
parents: 21915
diff changeset
   308
                    ^ (cat_lines o map string_of_thm) thms)
21120
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   309
          in (c, thms') end;
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   310
    fun rhs_of' thy (("", []), thms as [_]) =
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   311
          add_things_of thy (cons o snd) (NONE, thms) []
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   312
      | rhs_of' thy (c, thms) =
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   313
          add_things_of thy (cons o snd) (SOME c, thms) [];
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   314
    val (unif, _) =
21387
5d3d340cb783 clarified code for building function equation system; explicit check of type discipline
haftmann
parents: 21196
diff changeset
   315
      fold (fn (c, thms) => fold unify_const (rhs_of' thy (c, thms)))
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   316
        eqss' (Vartab.empty, maxidx);
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   317
    val eqss'' =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   318
      map (apply_unifier unif) eqss';
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   319
  in eqss'' end;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   320
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   321
fun merge_eqsyss thy raw_eqss funcgr =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   322
  let
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   323
    val eqss = specialize_typs thy funcgr raw_eqss;
20938
041badc7fcaf added keeping of funcgr
haftmann
parents: 20896
diff changeset
   324
    val tys = map (CodegenData.typ_funcs thy) eqss;
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   325
    val rhss = map (rhs_of thy) eqss;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   326
  in
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   327
    funcgr
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   328
    |> fold2 (fn (c, thms) => fn ty => Constgraph.new_node (c, (ty, thms))) eqss tys
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   329
    |> `(fn funcgr => map (insts_of_thms thy funcgr) eqss)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   330
    |-> (fn rhs_insts => fold2 (fn (c, _) => fn rhs_inst =>
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   331
          ensure_consts thy rhs_inst #> fold (curry Constgraph.add_edge c) rhs_inst) eqss rhs_insts)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   332
    |> fold2 (fn (c, _) => fn rhs => fold (curry Constgraph.add_edge c) rhs) eqss rhss
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   333
  end
20705
da71d46b8b2f fixed some mess
haftmann
parents: 20600
diff changeset
   334
and merge_new_eqsyss thy raw_eqss funcgr =
da71d46b8b2f fixed some mess
haftmann
parents: 20600
diff changeset
   335
  if exists (member CodegenConsts.eq_const (Constgraph.keys funcgr)) (map fst raw_eqss)
da71d46b8b2f fixed some mess
haftmann
parents: 20600
diff changeset
   336
  then funcgr
da71d46b8b2f fixed some mess
haftmann
parents: 20600
diff changeset
   337
  else merge_eqsyss thy raw_eqss funcgr
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   338
and ensure_consts thy cs funcgr =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   339
  fold (snd oo ensure_const thy funcgr) cs Constgraph.empty
20705
da71d46b8b2f fixed some mess
haftmann
parents: 20600
diff changeset
   340
  |> (fn auxgr => fold (merge_new_eqsyss thy)
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   341
       (map (AList.make (Constgraph.get_node auxgr))
21387
5d3d340cb783 clarified code for building function equation system; explicit check of type discipline
haftmann
parents: 21196
diff changeset
   342
       (rev (Constgraph.strong_conn auxgr))) funcgr)
5d3d340cb783 clarified code for building function equation system; explicit check of type discipline
haftmann
parents: 21196
diff changeset
   343
  handle INVALID (cs', msg) => raise INVALID (cs @ cs', msg);
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   344
21120
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   345
fun drop_classes thy tfrees thm =
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   346
  let
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   347
    val (_, thm') = Thm.varifyT' [] thm;
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   348
    val tvars = Term.add_tvars (Thm.prop_of thm') [];
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   349
    val unconstr = map (Thm.ctyp_of thy o TVar) tvars;
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   350
    val instmap = map2 (fn (v_i, _) => fn (v, sort) => pairself (Thm.ctyp_of thy)
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   351
      (TVar (v_i, []), TFree (v, sort))) tvars tfrees;
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   352
  in
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   353
    thm'
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   354
    |> fold Thm.unconstrainT unconstr
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   355
    |> Thm.instantiate (instmap, [])
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   356
    |> Tactic.rule_by_tactic ((REPEAT o CHANGED o ALLGOALS o Tactic.resolve_tac) (AxClass.class_intros thy))
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   357
  end;
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   358
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   359
in
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   360
21387
5d3d340cb783 clarified code for building function equation system; explicit check of type discipline
haftmann
parents: 21196
diff changeset
   361
val ensure_consts = (fn thy => fn cs => fn funcgr => ensure_consts thy
5d3d340cb783 clarified code for building function equation system; explicit check of type discipline
haftmann
parents: 21196
diff changeset
   362
  cs funcgr handle INVALID (cs', msg) => error (msg ^ ",\nwhile preprocessing equations for constant(s) "
5d3d340cb783 clarified code for building function equation system; explicit check of type discipline
haftmann
parents: 21196
diff changeset
   363
    ^ commas (map (CodegenConsts.string_of_const thy) cs')));
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   364
21120
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   365
fun make thy consts =
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   366
  Funcgr.change thy (ensure_consts thy consts);
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   367
21129
8e621232a865 clarified make_term interface
haftmann
parents: 21120
diff changeset
   368
fun make_term thy f ct =
21120
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   369
  let
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   370
    val _ = Sign.no_vars (Sign.pp thy) (Thm.term_of ct);
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   371
    val _ = Term.fold_types (Type.no_tvars #> K I) (Thm.term_of ct) ();
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   372
    val thm1 = CodegenData.preprocess_cterm thy ct;
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   373
    val ct' = Drule.dest_equals_rhs (Thm.cprop_of thm1);
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   374
    val consts = CodegenConsts.consts_of thy (Thm.term_of ct');
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   375
    val funcgr = make thy consts;
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   376
    val (_, thm2) = Thm.varifyT' [] thm1;
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   377
    val thm3 = Thm.reflexive (Drule.dest_equals_rhs (Thm.cprop_of thm2));
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   378
    val [(_, [thm4])] = specialize_typs thy funcgr [(("", []), [thm3])];
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   379
    val tfrees = Term.add_tfrees (Thm.prop_of thm1) [];
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   380
    fun inst thm =
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   381
      let
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   382
        val tvars = Term.add_tvars (Thm.prop_of thm) [];
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   383
        val instmap = map2 (fn (v_i, sort) => fn (v, _) => pairself (Thm.ctyp_of thy)
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   384
          (TVar (v_i, sort), TFree (v, sort))) tvars tfrees;
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   385
      in Thm.instantiate (instmap, []) thm end;
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   386
    val thm5 = inst thm2;
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   387
    val thm6 = inst thm4;
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   388
    val ct'' = Drule.dest_equals_rhs (Thm.cprop_of thm6);
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   389
    val cs = fold_aterms (fn Const c => cons c | _ => I) (Thm.term_of ct'') [];
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   390
    val drop = drop_classes thy tfrees;
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   391
    val funcgr' = ensure_consts thy
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   392
      (instdefs_of thy (fold (fold (insert (op =)) o insts_of thy funcgr) cs [])) funcgr
21129
8e621232a865 clarified make_term interface
haftmann
parents: 21120
diff changeset
   393
  in (f drop ct'' thm5, Funcgr.change thy (K funcgr')) end;
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   394
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   395
end; (*local*)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   396
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   397
fun print_funcgr thy funcgr =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   398
  AList.make (snd o Constgraph.get_node funcgr) (Constgraph.keys funcgr)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   399
  |> (map o apfst) (CodegenConsts.string_of_const thy)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   400
  |> sort (string_ord o pairself fst)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   401
  |> map (fn (s, thms) =>
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   402
       (Pretty.block o Pretty.fbreaks) (
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   403
         Pretty.str s
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   404
         :: map Display.pretty_thm thms
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   405
       ))
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   406
  |> Pretty.chunks
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   407
  |> Pretty.writeln;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   408
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   409
fun print_codethms thy consts =
21120
e333c844b057 refined algorithm
haftmann
parents: 20938
diff changeset
   410
  make thy consts |> print_funcgr thy;
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   411
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   412
fun print_codethms_e thy cs =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   413
  print_codethms thy (map (CodegenConsts.read_const thy) cs);
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   414
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   415
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   416
(** Isar **)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   417
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   418
structure P = OuterParse;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   419
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   420
val print_codethmsK = "print_codethms";
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   421
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   422
val print_codethmsP =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   423
  OuterSyntax.improper_command print_codethmsK "print code theorems of this theory" OuterKeyword.diag
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   424
    (Scan.option (P.$$$ "(" |-- Scan.repeat P.term --| P.$$$ ")")
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   425
      >> (fn NONE => CodegenData.print_thms
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   426
           | SOME cs => fn thy => print_codethms_e thy cs)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   427
      >> (fn f => Toplevel.no_timing o Toplevel.unknown_theory
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   428
      o Toplevel.keep (f o Toplevel.theory_of)));
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   429
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   430
val _ = OuterSyntax.add_parsers [print_codethmsP];
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   431
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   432
end; (*struct*)