src/Pure/Tools/codegen_funcgr.ML
author haftmann
Tue, 10 Oct 2006 09:17:24 +0200
changeset 20938 041badc7fcaf
parent 20896 1484c7af6d68
child 21120 e333c844b057
permissions -rw-r--r--
added keeping of funcgr
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;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    12
  val mk_funcgr: theory -> CodegenConsts.const list -> (string * typ) list -> T
20835
haftmann
parents: 20705
diff changeset
    13
  val all_deps_of: T -> CodegenConsts.const list -> CodegenConsts.const list list
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    14
  val get_funcs: T -> CodegenConsts.const -> thm list
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    15
  val get_func_typs: T -> (CodegenConsts.const * typ) list
20835
haftmann
parents: 20705
diff changeset
    16
  val normalize: theory -> thm list -> thm list
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    17
  val print_codethms: theory -> CodegenConsts.const list -> unit
20938
041badc7fcaf added keeping of funcgr
haftmann
parents: 20896
diff changeset
    18
  structure Constgraph : GRAPH
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    19
end;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    20
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    21
structure CodegenFuncgr: CODEGEN_FUNCGR =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    22
struct
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    23
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    24
(** code data **)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    25
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    26
structure Consttab = CodegenConsts.Consttab;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    27
structure Constgraph = GraphFun (
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    28
  type key = CodegenConsts.const;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    29
  val ord = CodegenConsts.const_ord;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    30
);
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    31
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    32
type T = (typ * thm list) Constgraph.T;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    33
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    34
structure Funcgr = CodeDataFun
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    35
(struct
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    36
  val name = "Pure/codegen_funcgr";
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    37
  type T = T;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    38
  val empty = Constgraph.empty;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    39
  fun merge _ _ = Constgraph.empty;
20938
041badc7fcaf added keeping of funcgr
haftmann
parents: 20896
diff changeset
    40
  fun purge _ NONE _ = Constgraph.empty
041badc7fcaf added keeping of funcgr
haftmann
parents: 20896
diff changeset
    41
    | purge _ (SOME cs) funcgr =
041badc7fcaf added keeping of funcgr
haftmann
parents: 20896
diff changeset
    42
        Constgraph.del_nodes ((Constgraph.all_succs funcgr 
041badc7fcaf added keeping of funcgr
haftmann
parents: 20896
diff changeset
    43
          o filter (can (Constgraph.get_node funcgr))) cs) funcgr;
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    44
end);
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    45
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    46
val _ = Context.add_setup Funcgr.init;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    47
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    48
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    49
(** theorem purification **)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    50
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    51
fun abs_norm thy thm =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    52
  let
20705
da71d46b8b2f fixed some mess
haftmann
parents: 20600
diff changeset
    53
    fun eta_expand thm =
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    54
      let
20705
da71d46b8b2f fixed some mess
haftmann
parents: 20600
diff changeset
    55
        val lhs = (fst o Logic.dest_equals o Drule.plain_prop_of) thm;
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    56
        val tys = (fst o strip_type o fastype_of) lhs;
20896
1484c7af6d68 cleaned up interfaces
haftmann
parents: 20855
diff changeset
    57
        val ctxt = Name.make_context (map (fst o fst) (Term.add_vars lhs []));
20705
da71d46b8b2f fixed some mess
haftmann
parents: 20600
diff changeset
    58
        val vs_ct = map (fn (v, ty) => cterm_of thy (Var ((v, 0), ty)))
20896
1484c7af6d68 cleaned up interfaces
haftmann
parents: 20855
diff changeset
    59
          (Name.names ctxt "a" tys);
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    60
      in
20705
da71d46b8b2f fixed some mess
haftmann
parents: 20600
diff changeset
    61
        fold (fn ct => fn thm => Thm.combination thm (Thm.reflexive ct)) vs_ct thm
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    62
      end;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    63
    fun beta_norm thm =
20705
da71d46b8b2f fixed some mess
haftmann
parents: 20600
diff changeset
    64
      let
da71d46b8b2f fixed some mess
haftmann
parents: 20600
diff changeset
    65
        val rhs = (snd o Logic.dest_equals o Drule.plain_prop_of) thm;
da71d46b8b2f fixed some mess
haftmann
parents: 20600
diff changeset
    66
        val thm' = Thm.beta_conversion true (cterm_of thy rhs);
da71d46b8b2f fixed some mess
haftmann
parents: 20600
diff changeset
    67
      in Thm.transitive thm thm' end;
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    68
  in
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    69
    thm
20705
da71d46b8b2f fixed some mess
haftmann
parents: 20600
diff changeset
    70
    |> eta_expand
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    71
    |> beta_norm
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    72
  end;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    73
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    74
fun canonical_tvars thy thm =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    75
  let
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    76
    fun mk_inst (v_i as (v, i), (v', sort)) (s as (maxidx, set, acc)) =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    77
      if v = v' orelse member (op =) set v then s
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    78
        else let
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    79
          val ty = TVar (v_i, sort)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    80
        in
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    81
          (maxidx + 1, v :: set,
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    82
            (ctyp_of thy ty, ctyp_of thy (TVar ((v', maxidx), sort))) :: acc)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    83
        end;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    84
    fun tvars_of thm = (fold_types o fold_atyps)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    85
      (fn TVar (v_i as (v, i), sort) => cons (v_i, (CodegenNames.purify_var v, sort))
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    86
        | _ => I) (prop_of thm) [];
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    87
    val maxidx = Thm.maxidx_of thm + 1;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    88
    val (_, _, inst) = fold mk_inst (tvars_of thm) (maxidx + 1, [], []);
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    89
  in Thm.instantiate (inst, []) thm end;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    90
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    91
fun canonical_vars thy thm =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    92
  let
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    93
    fun mk_inst (v_i as (v, i), (v', ty)) (s as (maxidx, set, acc)) =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    94
      if v = v' orelse member (op =) set v then s
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    95
        else let
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    96
          val t = if i = ~1 then Free (v, ty) else Var (v_i, ty)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    97
        in
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    98
          (maxidx + 1,  v :: set,
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
    99
            (cterm_of thy t, cterm_of thy (Var ((v', maxidx), ty))) :: acc)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   100
        end;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   101
    fun vars_of thm = fold_aterms
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   102
      (fn Var (v_i as (v, i), ty) => cons (v_i, (CodegenNames.purify_var v, ty))
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   103
        | _ => I) (prop_of thm) [];
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   104
    val maxidx = Thm.maxidx_of thm + 1;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   105
    val (_, _, inst) = fold mk_inst (vars_of thm) (maxidx + 1, [], []);
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   106
  in Thm.instantiate ([], inst) thm end;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   107
20835
haftmann
parents: 20705
diff changeset
   108
fun normalize thy thms =
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   109
  let
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   110
    fun burrow_thms f [] = []
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   111
      | burrow_thms f thms =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   112
          thms
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   113
          |> Conjunction.intr_list
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   114
          |> f
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   115
          |> Conjunction.elim_list;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   116
    fun unvarify thms =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   117
      #2 (#1 (Variable.import true thms (ProofContext.init thy)));
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   118
  in
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   119
    thms
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   120
    |> map (abs_norm thy)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   121
    |> burrow_thms (
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   122
        canonical_tvars thy
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   123
        #> canonical_vars thy
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   124
        #> Drule.zero_var_indexes
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   125
       )
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   126
  end;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   127
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   128
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   129
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   130
(** retrieval **)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   131
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   132
fun get_funcs funcgr (c_tys as (c, _)) =
20835
haftmann
parents: 20705
diff changeset
   133
  (these o Option.map snd o try (Constgraph.get_node funcgr)) c_tys;
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   134
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   135
fun get_func_typs funcgr =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   136
  AList.make (fst o Constgraph.get_node funcgr) (Constgraph.keys funcgr);
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   137
20835
haftmann
parents: 20705
diff changeset
   138
fun all_deps_of funcgr cs =
haftmann
parents: 20705
diff changeset
   139
  let
haftmann
parents: 20705
diff changeset
   140
    val conn = Constgraph.strong_conn funcgr;
haftmann
parents: 20705
diff changeset
   141
    val order = rev conn;
haftmann
parents: 20705
diff changeset
   142
  in
haftmann
parents: 20705
diff changeset
   143
    (map o filter) (member (op =) (Constgraph.all_succs funcgr cs)) order
haftmann
parents: 20705
diff changeset
   144
    |> filter_out null
haftmann
parents: 20705
diff changeset
   145
  end;
haftmann
parents: 20705
diff changeset
   146
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   147
local
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   148
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   149
fun add_things_of thy f (c, thms) =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   150
  (fold o fold_aterms)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   151
     (fn Const c_ty => let
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   152
            val c' = CodegenConsts.norm_of_typ thy c_ty
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   153
          in if CodegenConsts.eq_const (c, c') then I
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   154
          else f (c', c_ty) end
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   155
       | _ => I) (maps (op :: o swap o apfst (snd o strip_comb)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   156
            o Logic.dest_equals o Drule.plain_prop_of) thms)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   157
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   158
fun rhs_of thy (c, thms) =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   159
  Consttab.empty
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   160
  |> add_things_of thy (Consttab.update o rpair () o fst) (c, thms)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   161
  |> Consttab.keys;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   162
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   163
fun rhs_of' thy (c, thms) =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   164
  add_things_of thy (cons o snd) (c, thms) [];
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   165
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   166
fun insts_of thy funcgr (c, ty) =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   167
  let
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   168
    val tys = Sign.const_typargs thy (c, ty);
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   169
    val c' = CodegenConsts.norm thy (c, tys);
20705
da71d46b8b2f fixed some mess
haftmann
parents: 20600
diff changeset
   170
    val ty_decl = CodegenConsts.disc_typ_of_const thy
da71d46b8b2f fixed some mess
haftmann
parents: 20600
diff changeset
   171
      (fst o Constgraph.get_node funcgr o CodegenConsts.norm thy) (c, tys);
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   172
    val tys_decl = Sign.const_typargs thy (c, ty_decl);
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   173
    val pp = Sign.pp thy;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   174
    val algebra = Sign.classes_of thy;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   175
    fun classrel (x, _) _ = x;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   176
    fun constructor tyco xs class =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   177
      (tyco, class) :: maps (maps fst) xs;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   178
    fun variable (TVar (_, sort)) = map (pair []) sort
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   179
      | variable (TFree (_, sort)) = map (pair []) sort;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   180
    fun mk_inst ty (TVar (_, sort)) = cons (ty, sort)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   181
      | mk_inst ty (TFree (_, sort)) = cons (ty, sort)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   182
      | mk_inst (Type (tyco1, tys1)) (Type (tyco2, tys2)) =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   183
          if tyco1 <> tyco2 then error "bad instance"
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   184
          else fold2 mk_inst tys1 tys2;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   185
  in
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   186
    flat (maps (Sorts.of_sort_derivation pp algebra
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   187
      { classrel = classrel, constructor = constructor, variable = variable })
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   188
      (fold2 mk_inst tys tys_decl []))
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   189
  end;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   190
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   191
fun all_classops thy tyco class =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   192
  maps (AxClass.params_of thy)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   193
      (Graph.all_succs ((#classes o Sorts.rep_algebra o Sign.classes_of) thy) [class])
20705
da71d46b8b2f fixed some mess
haftmann
parents: 20600
diff changeset
   194
  |> AList.make (fn c => CodegenConsts.disc_typ_of_classop thy (c, [Type (tyco, [])]))
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   195
        (*typ_of_classop is very liberal in its type arguments*)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   196
  |> map (CodegenConsts.norm_of_typ thy);
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   197
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   198
fun instdefs_of thy insts =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   199
  let
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   200
    val thy_classes = (#classes o Sorts.rep_algebra o Sign.classes_of) thy;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   201
  in
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   202
    Symtab.empty
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   203
    |> fold (fn (tyco, class) =>
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   204
        Symtab.map_default (tyco, []) (insert (op =) class)) insts
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   205
    |> (fn tab => Symtab.fold (fn (tyco, classes) => append (maps (all_classops thy tyco)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   206
         (Graph.all_succs thy_classes classes))) tab [])
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 insts_of_thms thy funcgr c_thms =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   210
  let
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   211
    val insts = add_things_of thy (fn (_, c_ty) => fold (insert (op =))
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   212
      (insts_of thy funcgr c_ty)) c_thms [];
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   213
  in instdefs_of thy insts end;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   214
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   215
fun ensure_const thy funcgr c auxgr =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   216
  if can (Constgraph.get_node funcgr) c
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   217
    then (NONE, auxgr)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   218
  else if can (Constgraph.get_node auxgr) c
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   219
    then (SOME c, auxgr)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   220
  else if is_some (CodegenData.get_datatype_of_constr thy c) then
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   221
    auxgr
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   222
    |> Constgraph.new_node (c, [])
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   223
    |> pair (SOME c)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   224
  else let
20835
haftmann
parents: 20705
diff changeset
   225
    val thms = normalize thy (CodegenData.these_funcs thy c);
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   226
    val rhs = rhs_of thy (c, thms);
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   227
  in
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   228
    auxgr
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   229
    |> Constgraph.new_node (c, thms)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   230
    |> fold_map (ensure_const thy funcgr) rhs
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   231
    |-> (fn rhs' => fold (fn SOME c' => Constgraph.add_edge (c, c')
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   232
                           | NONE => I) rhs')
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   233
    |> pair (SOME c)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   234
  end;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   235
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   236
fun specialize_typs thy funcgr eqss =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   237
  let
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   238
    fun max k [] = k
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   239
      | max k (l::ls) = max (if k < l then l else k) ls;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   240
    fun typscheme_of (c, ty) =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   241
      try (Constgraph.get_node funcgr) (CodegenConsts.norm_of_typ thy (c, ty))
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   242
      |> Option.map fst;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   243
    fun incr_indices (c, thms) maxidx =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   244
      let
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   245
        val thms' = map (Thm.incr_indexes maxidx) thms;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   246
        val maxidx' = Int.max
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   247
          (maxidx, max ~1 (map Thm.maxidx_of thms') + 1);
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   248
      in ((c, thms'), maxidx') end;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   249
    val tsig = Sign.tsig_of thy;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   250
    fun unify_const thms (c, ty) (env, maxidx) =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   251
      case typscheme_of (c, ty)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   252
       of SOME ty_decl => let
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   253
            val ty_decl' = Logic.incr_tvar maxidx ty_decl;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   254
            val maxidx' = Int.max (Term.maxidx_of_typ ty_decl' + 1, maxidx);
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   255
          in Type.unify tsig (ty_decl', ty) (env, maxidx')
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   256
          handle TUNIFY => error ("Failed to instantiate\n"
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   257
            ^ (Sign.string_of_typ thy o Envir.norm_type env) ty_decl' ^ "\nto\n"
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   258
            ^ (Sign.string_of_typ thy o Envir.norm_type env) ty ^ ",\n"
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   259
            ^ "in function theorems\n"
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   260
            ^ cat_lines (map string_of_thm thms))
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   261
          end
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   262
        | NONE => (env, maxidx);
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   263
    fun apply_unifier unif (c, []) = (c, [])
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   264
      | apply_unifier unif (c, thms as thm :: _) =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   265
          let
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   266
            val ty = CodegenData.typ_func thy thm;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   267
            val ty' = Envir.norm_type unif ty;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   268
            val env = Type.typ_match (Sign.tsig_of thy) (ty, ty') Vartab.empty;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   269
            val inst = Thm.instantiate (Vartab.fold (fn (x_i, (sort, ty)) =>
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   270
              cons (Thm.ctyp_of thy (TVar (x_i, sort)), Thm.ctyp_of thy ty)) env [], []);
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   271
          in (c, map (Drule.zero_var_indexes o inst) thms) end;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   272
    val (eqss', maxidx) =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   273
      fold_map incr_indices eqss 0;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   274
    val (unif, _) =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   275
      fold (fn (c, thms) => fold (unify_const thms) (rhs_of' thy (c, thms)))
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   276
        eqss' (Vartab.empty, maxidx);
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   277
    val eqss'' =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   278
      map (apply_unifier unif) eqss';
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   279
  in eqss'' end;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   280
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   281
fun merge_eqsyss thy raw_eqss funcgr =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   282
  let
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   283
    val eqss = specialize_typs thy funcgr raw_eqss;
20938
041badc7fcaf added keeping of funcgr
haftmann
parents: 20896
diff changeset
   284
    val tys = map (CodegenData.typ_funcs thy) eqss;
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   285
    val rhss = map (rhs_of thy) eqss;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   286
  in
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   287
    funcgr
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   288
    |> fold2 (fn (c, thms) => fn ty => Constgraph.new_node (c, (ty, thms))) eqss tys
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   289
    |> `(fn funcgr => map (insts_of_thms thy funcgr) eqss)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   290
    |-> (fn rhs_insts => fold2 (fn (c, _) => fn rhs_inst =>
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   291
          ensure_consts thy rhs_inst #> fold (curry Constgraph.add_edge c) rhs_inst) eqss rhs_insts)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   292
    |> fold2 (fn (c, _) => fn rhs => fold (curry Constgraph.add_edge c) rhs) eqss rhss
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   293
  end
20705
da71d46b8b2f fixed some mess
haftmann
parents: 20600
diff changeset
   294
and merge_new_eqsyss thy raw_eqss funcgr =
da71d46b8b2f fixed some mess
haftmann
parents: 20600
diff changeset
   295
  if exists (member CodegenConsts.eq_const (Constgraph.keys funcgr)) (map fst raw_eqss)
da71d46b8b2f fixed some mess
haftmann
parents: 20600
diff changeset
   296
  then funcgr
da71d46b8b2f fixed some mess
haftmann
parents: 20600
diff changeset
   297
  else merge_eqsyss thy raw_eqss funcgr
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   298
and ensure_consts thy cs funcgr =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   299
  fold (snd oo ensure_const thy funcgr) cs Constgraph.empty
20705
da71d46b8b2f fixed some mess
haftmann
parents: 20600
diff changeset
   300
  |> (fn auxgr => fold (merge_new_eqsyss thy)
20600
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   301
       (map (AList.make (Constgraph.get_node auxgr))
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   302
       (rev (Constgraph.strong_conn auxgr))) funcgr);
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   303
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   304
in
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   305
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   306
val ensure_consts = ensure_consts;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   307
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   308
fun mk_funcgr thy consts cs =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   309
  Funcgr.change thy (
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   310
    ensure_consts thy consts
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   311
    #> (fn funcgr => ensure_consts thy
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   312
         (instdefs_of thy (fold (fold (insert (op =)) o insts_of thy funcgr) cs [])) funcgr)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   313
  );
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   314
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   315
end; (*local*)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   316
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   317
fun print_funcgr thy funcgr =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   318
  AList.make (snd o Constgraph.get_node funcgr) (Constgraph.keys funcgr)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   319
  |> (map o apfst) (CodegenConsts.string_of_const thy)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   320
  |> sort (string_ord o pairself fst)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   321
  |> map (fn (s, thms) =>
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   322
       (Pretty.block o Pretty.fbreaks) (
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   323
         Pretty.str s
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   324
         :: map Display.pretty_thm thms
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   325
       ))
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   326
  |> Pretty.chunks
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   327
  |> Pretty.writeln;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   328
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   329
fun print_codethms thy consts =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   330
  mk_funcgr thy consts [] |> print_funcgr thy;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   331
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   332
fun print_codethms_e thy cs =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   333
  print_codethms thy (map (CodegenConsts.read_const thy) cs);
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   334
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   335
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   336
(** Isar **)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   337
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   338
structure P = OuterParse;
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   339
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   340
val print_codethmsK = "print_codethms";
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   341
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   342
val print_codethmsP =
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   343
  OuterSyntax.improper_command print_codethmsK "print code theorems of this theory" OuterKeyword.diag
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   344
    (Scan.option (P.$$$ "(" |-- Scan.repeat P.term --| P.$$$ ")")
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   345
      >> (fn NONE => CodegenData.print_thms
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   346
           | SOME cs => fn thy => print_codethms_e thy cs)
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   347
      >> (fn f => Toplevel.no_timing o Toplevel.unknown_theory
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   348
      o Toplevel.keep (f o Toplevel.theory_of)));
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   349
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   350
val _ = OuterSyntax.add_parsers [print_codethmsP];
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   351
6d75e02ed285 added codegen_data
haftmann
parents:
diff changeset
   352
end; (*struct*)