src/Pure/Tools/nbe.ML
author wenzelm
Wed, 25 Jul 2007 17:05:49 +0200
changeset 23981 03b71bf91318
parent 23934 79393cb9c0a6
child 24219 e558fe311376
permissions -rw-r--r--
added trace flag, official tracing operation; added named CRITICAL'; schedule: tuned signature, actually observe dependencies on running tasks;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
20856
9f7f0bf89e7d cleaned up some mess
haftmann
parents: 20846
diff changeset
     1
(*  Title:      Pure/nbe.ML
9f7f0bf89e7d cleaned up some mess
haftmann
parents: 20846
diff changeset
     2
    ID:         $Id$
19147
0f584853b6a4 added nbe, updated neb_*
nipkow
parents:
diff changeset
     3
    Author:     Tobias Nipkow, Florian Haftmann, TU Muenchen
0f584853b6a4 added nbe, updated neb_*
nipkow
parents:
diff changeset
     4
19177
68c6824d8bb6 improvements for nbe
haftmann
parents: 19167
diff changeset
     5
Toplevel theory interface for "normalization by evaluation"
19147
0f584853b6a4 added nbe, updated neb_*
nipkow
parents:
diff changeset
     6
*)
0f584853b6a4 added nbe, updated neb_*
nipkow
parents:
diff changeset
     7
19177
68c6824d8bb6 improvements for nbe
haftmann
parents: 19167
diff changeset
     8
signature NBE =
19147
0f584853b6a4 added nbe, updated neb_*
nipkow
parents:
diff changeset
     9
sig
20856
9f7f0bf89e7d cleaned up some mess
haftmann
parents: 20846
diff changeset
    10
  (*preconditions: no Vars/TVars in term*)
20602
96fa2cf465f5 moved part of normalization oracle here
haftmann
parents: 20191
diff changeset
    11
  val normalization_conv: cterm -> thm
19962
016ba2d907a7 new function norm_term
nipkow
parents: 19830
diff changeset
    12
  val lookup: string -> NBE_Eval.Univ
19795
746274ca400b added type inference at the end of normalization
nipkow
parents: 19341
diff changeset
    13
  val update: string * NBE_Eval.Univ -> unit
20846
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
    14
  val trace: bool ref
19147
0f584853b6a4 added nbe, updated neb_*
nipkow
parents:
diff changeset
    15
end;
0f584853b6a4 added nbe, updated neb_*
nipkow
parents:
diff changeset
    16
19177
68c6824d8bb6 improvements for nbe
haftmann
parents: 19167
diff changeset
    17
structure NBE: NBE =
19147
0f584853b6a4 added nbe, updated neb_*
nipkow
parents:
diff changeset
    18
struct
0f584853b6a4 added nbe, updated neb_*
nipkow
parents:
diff changeset
    19
20846
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
    20
val trace = ref false;
21124
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
    21
fun tracing f x = if !trace then (Output.tracing (f x); x) else x;
19968
2180f0f443af fixed stale theory bug
haftmann
parents: 19962
diff changeset
    22
21068
a6f47c0e7dbb cleanup
haftmann
parents: 21002
diff changeset
    23
a6f47c0e7dbb cleanup
haftmann
parents: 21002
diff changeset
    24
(** data setup **)
a6f47c0e7dbb cleanup
haftmann
parents: 21002
diff changeset
    25
a6f47c0e7dbb cleanup
haftmann
parents: 21002
diff changeset
    26
(* preproc and postproc attributes *)
19968
2180f0f443af fixed stale theory bug
haftmann
parents: 19962
diff changeset
    27
20920
07f279940664 added pre/post-processor equations
nipkow
parents: 20856
diff changeset
    28
structure NBE_Rewrite = TheoryDataFun
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22728
diff changeset
    29
(
20920
07f279940664 added pre/post-processor equations
nipkow
parents: 20856
diff changeset
    30
  type T = thm list * thm list
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22728
diff changeset
    31
  val empty = ([], []);
20920
07f279940664 added pre/post-processor equations
nipkow
parents: 20856
diff changeset
    32
  val copy = I;
07f279940664 added pre/post-processor equations
nipkow
parents: 20856
diff changeset
    33
  val extend = I;
07f279940664 added pre/post-processor equations
nipkow
parents: 20856
diff changeset
    34
  fun merge _ ((pres1,posts1), (pres2,posts2)) =
22360
26ead7ed4f4b moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
wenzelm
parents: 22213
diff changeset
    35
    (Library.merge Thm.eq_thm (pres1,pres2), Library.merge Thm.eq_thm (posts1,posts2))
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22728
diff changeset
    36
);
20920
07f279940664 added pre/post-processor equations
nipkow
parents: 20856
diff changeset
    37
21068
a6f47c0e7dbb cleanup
haftmann
parents: 21002
diff changeset
    38
val _ =
a6f47c0e7dbb cleanup
haftmann
parents: 21002
diff changeset
    39
  let
a6f47c0e7dbb cleanup
haftmann
parents: 21002
diff changeset
    40
    fun map_data f = Context.mapping (NBE_Rewrite.map f) I;
a6f47c0e7dbb cleanup
haftmann
parents: 21002
diff changeset
    41
    fun attr_pre (thy, thm) =
22360
26ead7ed4f4b moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
wenzelm
parents: 22213
diff changeset
    42
      ((map_data o apfst) (insert Thm.eq_thm thm) thy, thm)
21068
a6f47c0e7dbb cleanup
haftmann
parents: 21002
diff changeset
    43
    fun attr_post (thy, thm) = 
22360
26ead7ed4f4b moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
wenzelm
parents: 22213
diff changeset
    44
      ((map_data o apsnd) (insert Thm.eq_thm thm) thy, thm)
21068
a6f47c0e7dbb cleanup
haftmann
parents: 21002
diff changeset
    45
    val attr = Attrib.syntax (Scan.lift (Args.$$$ "pre" >> K attr_pre
a6f47c0e7dbb cleanup
haftmann
parents: 21002
diff changeset
    46
      || Args.$$$ "post" >> K attr_post));
a6f47c0e7dbb cleanup
haftmann
parents: 21002
diff changeset
    47
  in
a6f47c0e7dbb cleanup
haftmann
parents: 21002
diff changeset
    48
    Context.add_setup (
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22728
diff changeset
    49
      Attrib.add_attributes
21068
a6f47c0e7dbb cleanup
haftmann
parents: 21002
diff changeset
    50
          [("normal", attr, "declare rewrite theorems for normalization")]
a6f47c0e7dbb cleanup
haftmann
parents: 21002
diff changeset
    51
    )
a6f47c0e7dbb cleanup
haftmann
parents: 21002
diff changeset
    52
  end;
20920
07f279940664 added pre/post-processor equations
nipkow
parents: 20856
diff changeset
    53
22213
2dd23002c465 adjusted to new codegen_funcgr interface
haftmann
parents: 22144
diff changeset
    54
fun the_pres thy =
21068
a6f47c0e7dbb cleanup
haftmann
parents: 21002
diff changeset
    55
  let
21506
b2a673894ce5 prefer Proof.context over Context.generic;
wenzelm
parents: 21129
diff changeset
    56
    val ctxt = ProofContext.init thy;
b2a673894ce5 prefer Proof.context over Context.generic;
wenzelm
parents: 21129
diff changeset
    57
    val pres = (map (LocalDefs.meta_rewrite_rule ctxt) o fst) (NBE_Rewrite.get thy)
22213
2dd23002c465 adjusted to new codegen_funcgr interface
haftmann
parents: 22144
diff changeset
    58
  in pres end
20920
07f279940664 added pre/post-processor equations
nipkow
parents: 20856
diff changeset
    59
07f279940664 added pre/post-processor equations
nipkow
parents: 20856
diff changeset
    60
fun apply_posts thy =
21068
a6f47c0e7dbb cleanup
haftmann
parents: 21002
diff changeset
    61
  let
21506
b2a673894ce5 prefer Proof.context over Context.generic;
wenzelm
parents: 21129
diff changeset
    62
    val ctxt = ProofContext.init thy;
b2a673894ce5 prefer Proof.context over Context.generic;
wenzelm
parents: 21129
diff changeset
    63
    val posts = (map (LocalDefs.meta_rewrite_rule ctxt) o snd) (NBE_Rewrite.get thy)
21708
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21506
diff changeset
    64
  in MetaSimplifier.rewrite false posts end
20920
07f279940664 added pre/post-processor equations
nipkow
parents: 20856
diff changeset
    65
22213
2dd23002c465 adjusted to new codegen_funcgr interface
haftmann
parents: 22144
diff changeset
    66
(* theorem store *)
2dd23002c465 adjusted to new codegen_funcgr interface
haftmann
parents: 22144
diff changeset
    67
23136
5a0378eada70 simplified data setup
haftmann
parents: 23090
diff changeset
    68
structure Funcgr = CodegenFuncgrRetrieval (val rewrites = the_pres);
20920
07f279940664 added pre/post-processor equations
nipkow
parents: 20856
diff changeset
    69
21068
a6f47c0e7dbb cleanup
haftmann
parents: 21002
diff changeset
    70
(* code store *)
a6f47c0e7dbb cleanup
haftmann
parents: 21002
diff changeset
    71
20846
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
    72
structure NBE_Data = CodeDataFun
19147
0f584853b6a4 added nbe, updated neb_*
nipkow
parents:
diff changeset
    73
(struct
22213
2dd23002c465 adjusted to new codegen_funcgr interface
haftmann
parents: 22144
diff changeset
    74
  type T = NBE_Eval.Univ Symtab.table;
2dd23002c465 adjusted to new codegen_funcgr interface
haftmann
parents: 22144
diff changeset
    75
  val empty = Symtab.empty;
2dd23002c465 adjusted to new codegen_funcgr interface
haftmann
parents: 22144
diff changeset
    76
  fun merge _ = Symtab.merge (K true);
2dd23002c465 adjusted to new codegen_funcgr interface
haftmann
parents: 22144
diff changeset
    77
  fun purge _ _ _ = Symtab.empty;
19147
0f584853b6a4 added nbe, updated neb_*
nipkow
parents:
diff changeset
    78
end);
0f584853b6a4 added nbe, updated neb_*
nipkow
parents:
diff changeset
    79
19795
746274ca400b added type inference at the end of normalization
nipkow
parents: 19341
diff changeset
    80
21124
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
    81
(** norm by eval **)
21068
a6f47c0e7dbb cleanup
haftmann
parents: 21002
diff changeset
    82
19968
2180f0f443af fixed stale theory bug
haftmann
parents: 19962
diff changeset
    83
(* sandbox communication *)
19147
0f584853b6a4 added nbe, updated neb_*
nipkow
parents:
diff changeset
    84
19177
68c6824d8bb6 improvements for nbe
haftmann
parents: 19167
diff changeset
    85
val tab : NBE_Eval.Univ Symtab.table ref = ref Symtab.empty;
68c6824d8bb6 improvements for nbe
haftmann
parents: 19167
diff changeset
    86
fun lookup s = (the o Symtab.lookup (!tab)) s;
68c6824d8bb6 improvements for nbe
haftmann
parents: 19167
diff changeset
    87
fun update sx = (tab := Symtab.update sx (!tab));
68c6824d8bb6 improvements for nbe
haftmann
parents: 19167
diff changeset
    88
20846
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
    89
local
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
    90
21124
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
    91
(* function generation *)
20846
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
    92
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
    93
fun generate thy funs =
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
    94
  let
21124
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
    95
    (* FIXME better turn this into a function
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
    96
        NBE_Eval.Univ Symtab.table -> NBE_Eval.Univ Symtab.table
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
    97
        with implicit side effect *)
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
    98
    fun use_code NONE = ()
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
    99
      | use_code (SOME s) =
21991
04528ce9ded5 fixed output
haftmann
parents: 21708
diff changeset
   100
          (tracing (fn () => "\n---generated code:\n" ^ s) ();
22144
c33450acd873 use_text: added name argument;
wenzelm
parents: 22033
diff changeset
   101
           use_text "" (Output.tracing o enclose "\n---compiler echo:\n" "\n---\n",
21124
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   102
                Output.tracing o enclose "\n--- compiler echo (with error!):\n" "\n---\n")
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   103
            (!trace) s);
21991
04528ce9ded5 fixed output
haftmann
parents: 21708
diff changeset
   104
    val _ = tracing (fn () => "new definitions: " ^ (commas o maps (map fst)) funs) ();
20846
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
   105
    val _ = tab := NBE_Data.get thy;;
23226
441f8a0bd766 removed obsolete Library.seq;
wenzelm
parents: 23136
diff changeset
   106
    val _ = List.app (use_code o NBE_Codegen.generate thy
20846
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
   107
      (fn s => Symtab.defined (!tab) s)) funs;
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
   108
  in NBE_Data.change thy (K (!tab)) end;
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
   109
22213
2dd23002c465 adjusted to new codegen_funcgr interface
haftmann
parents: 22144
diff changeset
   110
fun ensure_funs thy funcgr t =
20846
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
   111
  let
22554
d1499fff65d8 simplified constant representation in code generator
haftmann
parents: 22360
diff changeset
   112
    fun consts_of thy t =
d1499fff65d8 simplified constant representation in code generator
haftmann
parents: 22360
diff changeset
   113
      fold_aterms (fn Const c => cons (CodegenConsts.const_of_cexpr thy c) | _ => I) t []
d1499fff65d8 simplified constant representation in code generator
haftmann
parents: 22360
diff changeset
   114
    val consts = consts_of thy t;
20846
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
   115
    val nbe_tab = NBE_Data.get thy;
22213
2dd23002c465 adjusted to new codegen_funcgr interface
haftmann
parents: 22144
diff changeset
   116
  in
2dd23002c465 adjusted to new codegen_funcgr interface
haftmann
parents: 22144
diff changeset
   117
    CodegenFuncgr.deps funcgr consts
2dd23002c465 adjusted to new codegen_funcgr interface
haftmann
parents: 22144
diff changeset
   118
    |> (map o filter_out) (Symtab.defined nbe_tab o CodegenNames.const thy)
2dd23002c465 adjusted to new codegen_funcgr interface
haftmann
parents: 22144
diff changeset
   119
    |> filter_out null
2dd23002c465 adjusted to new codegen_funcgr interface
haftmann
parents: 22144
diff changeset
   120
    |> (map o map) (fn c => (CodegenNames.const thy c, CodegenFuncgr.funcs funcgr c))
2dd23002c465 adjusted to new codegen_funcgr interface
haftmann
parents: 22144
diff changeset
   121
    |> generate thy
2dd23002c465 adjusted to new codegen_funcgr interface
haftmann
parents: 22144
diff changeset
   122
  end;
21124
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   123
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   124
(* term evaluation *)
20846
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
   125
22213
2dd23002c465 adjusted to new codegen_funcgr interface
haftmann
parents: 22144
diff changeset
   126
fun eval_term thy funcgr t =
21124
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   127
  let
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   128
    fun subst_Frees [] = I
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   129
      | subst_Frees inst =
23090
eb3000a5c478 fixes tvar issue in type inference
haftmann
parents: 22902
diff changeset
   130
          Term.map_aterms (fn (t as Free (s, _)) => the_default t (AList.lookup (op =) inst s)
eb3000a5c478 fixes tvar issue in type inference
haftmann
parents: 22902
diff changeset
   131
                            | t => t);
21124
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   132
    val anno_vars =
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   133
      subst_Frees (map (fn (s, T) => (s, Free (s, T))) (Term.add_frees t []))
23090
eb3000a5c478 fixes tvar issue in type inference
haftmann
parents: 22902
diff changeset
   134
      #> subst_Vars (map (fn (ixn, T) => (ixn, Var (ixn, T))) (Term.add_vars t []))
21124
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   135
    fun check_tvars t = if null (Term.term_tvars t) then t else
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   136
      error ("Illegal schematic type variables in normalized term: "
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   137
        ^ setmp show_types true (Sign.string_of_term thy) t);
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   138
    val ty = type_of t;
22728
ecbbdf50df2f simplified ProofContext.infer_types(_pats);
wenzelm
parents: 22692
diff changeset
   139
    fun constrain t =
23090
eb3000a5c478 fixes tvar issue in type inference
haftmann
parents: 22902
diff changeset
   140
      singleton (ProofContext.infer_types_pats (ProofContext.init thy)) (TypeInfer.constrain t ty);
22213
2dd23002c465 adjusted to new codegen_funcgr interface
haftmann
parents: 22144
diff changeset
   141
    val _ = ensure_funs thy funcgr t;
21124
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   142
  in
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   143
    t
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   144
    |> tracing (fn t => "Input:\n" ^ Display.raw_string_of_term t)
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   145
    |> NBE_Eval.eval thy (!tab)
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   146
    |> tracing (fn nt => "Normalized:\n" ^ NBE_Eval.string_of_nterm nt)
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   147
    |> NBE_Codegen.nterm_to_term thy
23090
eb3000a5c478 fixes tvar issue in type inference
haftmann
parents: 22902
diff changeset
   148
    |> tracing (fn t => "Converted back:\n" ^ setmp show_types true Display.raw_string_of_term t)
eb3000a5c478 fixes tvar issue in type inference
haftmann
parents: 22902
diff changeset
   149
    |> tracing (fn _ => "Term type:\n" ^ Display.raw_string_of_typ ty)
21124
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   150
    |> anno_vars
23090
eb3000a5c478 fixes tvar issue in type inference
haftmann
parents: 22902
diff changeset
   151
    |> tracing (fn t => "Vars typed:\n" ^ setmp show_types true Display.raw_string_of_term t)
21124
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   152
    |> constrain
23090
eb3000a5c478 fixes tvar issue in type inference
haftmann
parents: 22902
diff changeset
   153
    |> tracing (fn t => "Types inferred:\n" ^ setmp show_types true Display.raw_string_of_term t)
21124
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   154
    |> check_tvars
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   155
  end;
19795
746274ca400b added type inference at the end of normalization
nipkow
parents: 19341
diff changeset
   156
21124
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   157
(* evaluation oracle *)
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   158
22213
2dd23002c465 adjusted to new codegen_funcgr interface
haftmann
parents: 22144
diff changeset
   159
exception Normalization of CodegenFuncgr.T * term;
19795
746274ca400b added type inference at the end of normalization
nipkow
parents: 19341
diff changeset
   160
22213
2dd23002c465 adjusted to new codegen_funcgr interface
haftmann
parents: 22144
diff changeset
   161
fun normalization_oracle (thy, Normalization (funcgr, t)) =
2dd23002c465 adjusted to new codegen_funcgr interface
haftmann
parents: 22144
diff changeset
   162
  Logic.mk_equals (t, eval_term thy funcgr t);
21124
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   163
22213
2dd23002c465 adjusted to new codegen_funcgr interface
haftmann
parents: 22144
diff changeset
   164
fun normalization_invoke thy funcgr t =
2dd23002c465 adjusted to new codegen_funcgr interface
haftmann
parents: 22144
diff changeset
   165
  Thm.invoke_oracle_i thy "Pure.normalization" (thy, Normalization (funcgr, t));
19968
2180f0f443af fixed stale theory bug
haftmann
parents: 19962
diff changeset
   166
20846
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
   167
in
19968
2180f0f443af fixed stale theory bug
haftmann
parents: 19962
diff changeset
   168
21124
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   169
(* interface *)
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   170
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   171
fun normalization_conv ct =
20846
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
   172
  let
21124
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   173
    val thy = Thm.theory_of_cterm ct;
22213
2dd23002c465 adjusted to new codegen_funcgr interface
haftmann
parents: 22144
diff changeset
   174
    fun mk funcgr drop_classes ct thm1 =
21129
8e621232a865 clarified make_term interface
haftmann
parents: 21124
diff changeset
   175
      let
8e621232a865 clarified make_term interface
haftmann
parents: 21124
diff changeset
   176
        val t = Thm.term_of ct;
22213
2dd23002c465 adjusted to new codegen_funcgr interface
haftmann
parents: 22144
diff changeset
   177
        val thm2 = normalization_invoke thy funcgr t;
22902
ac833b4bb7ee moved some Drule operations to Thm (see more_thm.ML);
wenzelm
parents: 22846
diff changeset
   178
        val thm3 = apply_posts thy (Thm.rhs_of thm2);
21129
8e621232a865 clarified make_term interface
haftmann
parents: 21124
diff changeset
   179
        val thm23 = drop_classes (Thm.transitive thm2 thm3);
8e621232a865 clarified make_term interface
haftmann
parents: 21124
diff changeset
   180
      in
8e621232a865 clarified make_term interface
haftmann
parents: 21124
diff changeset
   181
        Thm.transitive thm1 thm23 handle THM _ =>
8e621232a865 clarified make_term interface
haftmann
parents: 21124
diff changeset
   182
          error ("normalization_conv - could not construct proof:\n"
8e621232a865 clarified make_term interface
haftmann
parents: 21124
diff changeset
   183
          ^ (cat_lines o map string_of_thm) [thm1, thm2, thm3])
8e621232a865 clarified make_term interface
haftmann
parents: 21124
diff changeset
   184
      end;
22213
2dd23002c465 adjusted to new codegen_funcgr interface
haftmann
parents: 22144
diff changeset
   185
  in fst (Funcgr.make_term thy mk ct) end;
20846
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
   186
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
   187
fun norm_print_term ctxt modes t =
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
   188
  let
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
   189
    val thy = ProofContext.theory_of ctxt;
21124
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   190
    val ct = Thm.cterm_of thy t;
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   191
    val (_, t') = (Logic.dest_equals o Thm.prop_of o normalization_conv) ct;
20846
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
   192
    val ty = Term.type_of t';
23934
79393cb9c0a6 added with_modes, with_default;
wenzelm
parents: 23226
diff changeset
   193
    val p = PrintMode.with_modes modes (fn () =>
20846
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
   194
      Pretty.block [Pretty.quote (ProofContext.pretty_term ctxt t'), Pretty.fbrk,
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
   195
        Pretty.str "::", Pretty.brk 1, Pretty.quote (ProofContext.pretty_typ ctxt ty)]) ();
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
   196
  in Pretty.writeln p end;
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
   197
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
   198
fun norm_print_term_e (modes, raw_t) state =
21506
b2a673894ce5 prefer Proof.context over Context.generic;
wenzelm
parents: 21129
diff changeset
   199
  let val ctxt = Toplevel.context_of state
20846
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
   200
  in norm_print_term ctxt modes (ProofContext.read_term ctxt raw_t) end;
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
   201
20602
96fa2cf465f5 moved part of normalization oracle here
haftmann
parents: 20191
diff changeset
   202
val _ = Context.add_setup
20677
5ce43b38a175 tuned oracle name;
wenzelm
parents: 20602
diff changeset
   203
  (Theory.add_oracle ("normalization", normalization_oracle));
19177
68c6824d8bb6 improvements for nbe
haftmann
parents: 19167
diff changeset
   204
21124
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   205
end; (*local*)
8648b5dd6a87 constructing proof
haftmann
parents: 21068
diff changeset
   206
20846
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
   207
19968
2180f0f443af fixed stale theory bug
haftmann
parents: 19962
diff changeset
   208
(* Isar setup *)
2180f0f443af fixed stale theory bug
haftmann
parents: 19962
diff changeset
   209
2180f0f443af fixed stale theory bug
haftmann
parents: 19962
diff changeset
   210
local structure P = OuterParse and K = OuterKeyword in
19147
0f584853b6a4 added nbe, updated neb_*
nipkow
parents:
diff changeset
   211
20846
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
   212
val opt_modes = Scan.optional (P.$$$ "(" |-- P.!!! (Scan.repeat1 P.xname --| P.$$$ ")")) [];
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
   213
19147
0f584853b6a4 added nbe, updated neb_*
nipkow
parents:
diff changeset
   214
val nbeP =
20846
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
   215
  OuterSyntax.improper_command "normal_form" "normalize term by evaluation" K.diag
5fde744176d7 various code refinements
haftmann
parents: 20677
diff changeset
   216
    (opt_modes -- P.typ >> (Toplevel.keep o norm_print_term_e));
19968
2180f0f443af fixed stale theory bug
haftmann
parents: 19962
diff changeset
   217
2180f0f443af fixed stale theory bug
haftmann
parents: 19962
diff changeset
   218
end;
19147
0f584853b6a4 added nbe, updated neb_*
nipkow
parents:
diff changeset
   219
0f584853b6a4 added nbe, updated neb_*
nipkow
parents:
diff changeset
   220
val _ = OuterSyntax.add_parsers [nbeP];
0f584853b6a4 added nbe, updated neb_*
nipkow
parents:
diff changeset
   221
19177
68c6824d8bb6 improvements for nbe
haftmann
parents: 19167
diff changeset
   222
end;