src/Pure/ML/ml_thms.ML
author Fabian Huch <huch@in.tum.de>
Fri, 23 Aug 2024 15:30:09 +0200
changeset 80949 97924a26a5c3
parent 78812 d769a183d51d
permissions -rw-r--r--
add comments to rendering, e.g. to collect them from build database;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
27389
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/ML/ml_thms.ML
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
     3
45592
8baa0b7f3f66 added ML antiquotation @{attributes};
wenzelm
parents: 45198
diff changeset
     4
Attribute source and theorem values within ML.
27389
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
     5
*)
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
     6
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
     7
signature ML_THMS =
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
     8
sig
58011
bc6bced136e5 tuned signature -- moved type src to Token, without aliases;
wenzelm
parents: 56304
diff changeset
     9
  val the_attributes: Proof.context -> int -> Token.src list
48782
e955964d89cb more direct embedding of abstract thm values into the ML environment -- avoidance of repeated ML_Thms.the_thm(s) considerably reduces compilation time for Poly/ML 5.4.x;
wenzelm
parents: 48776
diff changeset
    10
  val the_thmss: Proof.context -> thm list list
78812
d769a183d51d simprocs may be distinguished via 'identifier': only works for ML antiquotation (see also 13252110a6fe);
wenzelm
parents: 78795
diff changeset
    11
  val thm_binding: string -> bool -> thm list -> Proof.context ->
d769a183d51d simprocs may be distinguished via 'identifier': only works for ML antiquotation (see also 13252110a6fe);
wenzelm
parents: 78795
diff changeset
    12
    (Proof.context -> string * string) * Proof.context
74606
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74604
diff changeset
    13
  val embedded_lemma: (Proof.context -> thm list * (term list * Proof.context)) parser
56199
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 56072
diff changeset
    14
  val get_stored_thms: unit -> thm list
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 56072
diff changeset
    15
  val get_stored_thm: unit -> thm
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 56072
diff changeset
    16
  val store_thms: string * thm list -> unit
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 56072
diff changeset
    17
  val store_thm: string * thm -> unit
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 56072
diff changeset
    18
  val bind_thm: string * thm -> unit
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 56072
diff changeset
    19
  val bind_thms: string * thm list -> unit
27389
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    20
end;
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    21
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    22
structure ML_Thms: ML_THMS =
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    23
struct
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    24
45592
8baa0b7f3f66 added ML antiquotation @{attributes};
wenzelm
parents: 45198
diff changeset
    25
(* auxiliary data *)
27389
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    26
48782
e955964d89cb more direct embedding of abstract thm values into the ML environment -- avoidance of repeated ML_Thms.the_thm(s) considerably reduces compilation time for Poly/ML 5.4.x;
wenzelm
parents: 48776
diff changeset
    27
type thms = (string * bool) * thm list;  (*name, single, value*)
e955964d89cb more direct embedding of abstract thm values into the ML environment -- avoidance of repeated ML_Thms.the_thm(s) considerably reduces compilation time for Poly/ML 5.4.x;
wenzelm
parents: 48776
diff changeset
    28
45592
8baa0b7f3f66 added ML antiquotation @{attributes};
wenzelm
parents: 45198
diff changeset
    29
structure Data = Proof_Data
27389
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    30
(
58011
bc6bced136e5 tuned signature -- moved type src to Token, without aliases;
wenzelm
parents: 56304
diff changeset
    31
  type T = Token.src list Inttab.table * thms list;
48782
e955964d89cb more direct embedding of abstract thm values into the ML environment -- avoidance of repeated ML_Thms.the_thm(s) considerably reduces compilation time for Poly/ML 5.4.x;
wenzelm
parents: 48776
diff changeset
    32
  fun init _ = (Inttab.empty, []);
27389
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    33
);
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    34
45592
8baa0b7f3f66 added ML antiquotation @{attributes};
wenzelm
parents: 45198
diff changeset
    35
val put_attributes = Data.map o apfst o Inttab.update;
8baa0b7f3f66 added ML antiquotation @{attributes};
wenzelm
parents: 45198
diff changeset
    36
fun the_attributes ctxt name = the (Inttab.lookup (fst (Data.get ctxt)) name);
8baa0b7f3f66 added ML antiquotation @{attributes};
wenzelm
parents: 45198
diff changeset
    37
48782
e955964d89cb more direct embedding of abstract thm values into the ML environment -- avoidance of repeated ML_Thms.the_thm(s) considerably reduces compilation time for Poly/ML 5.4.x;
wenzelm
parents: 48776
diff changeset
    38
val get_thmss = snd o Data.get;
e955964d89cb more direct embedding of abstract thm values into the ML environment -- avoidance of repeated ML_Thms.the_thm(s) considerably reduces compilation time for Poly/ML 5.4.x;
wenzelm
parents: 48776
diff changeset
    39
val the_thmss = map snd o get_thmss;
e955964d89cb more direct embedding of abstract thm values into the ML environment -- avoidance of repeated ML_Thms.the_thm(s) considerably reduces compilation time for Poly/ML 5.4.x;
wenzelm
parents: 48776
diff changeset
    40
val cons_thms = Data.map o apsnd o cons;
45592
8baa0b7f3f66 added ML antiquotation @{attributes};
wenzelm
parents: 45198
diff changeset
    41
8baa0b7f3f66 added ML antiquotation @{attributes};
wenzelm
parents: 45198
diff changeset
    42
8baa0b7f3f66 added ML antiquotation @{attributes};
wenzelm
parents: 45198
diff changeset
    43
(* attribute source *)
27389
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    44
53171
a5e54d4d9081 added Theory.setup convenience;
wenzelm
parents: 53169
diff changeset
    45
val _ = Theory.setup
67147
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 66067
diff changeset
    46
  (ML_Antiquotation.declaration \<^binding>\<open>attributes\<close> Attrib.attribs
61814
1ca1142e1711 clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents: 59127
diff changeset
    47
    (fn _ => fn srcs => fn ctxt =>
59112
e670969f34df expand ML cartouches to Input.source;
wenzelm
parents: 58866
diff changeset
    48
      let val i = serial () in
e670969f34df expand ML cartouches to Input.source;
wenzelm
parents: 58866
diff changeset
    49
        ctxt
61814
1ca1142e1711 clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents: 59127
diff changeset
    50
        |> put_attributes (i, srcs)
73549
a2c589d5e1e4 clarified signature: more detailed token positions for antiquotations;
wenzelm
parents: 72054
diff changeset
    51
        |> ML_Antiquotation.value_decl "attributes"
59112
e670969f34df expand ML cartouches to Input.source;
wenzelm
parents: 58866
diff changeset
    52
            ("ML_Thms.the_attributes ML_context " ^ string_of_int i)
e670969f34df expand ML cartouches to Input.source;
wenzelm
parents: 58866
diff changeset
    53
      end));
45592
8baa0b7f3f66 added ML antiquotation @{attributes};
wenzelm
parents: 45198
diff changeset
    54
8baa0b7f3f66 added ML antiquotation @{attributes};
wenzelm
parents: 45198
diff changeset
    55
8baa0b7f3f66 added ML antiquotation @{attributes};
wenzelm
parents: 45198
diff changeset
    56
(* fact references *)
27389
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    57
56069
451d5b73f8cf simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents: 55997
diff changeset
    58
fun thm_binding kind is_single thms ctxt =
48782
e955964d89cb more direct embedding of abstract thm values into the ML environment -- avoidance of repeated ML_Thms.the_thm(s) considerably reduces compilation time for Poly/ML 5.4.x;
wenzelm
parents: 48776
diff changeset
    59
  let
53169
e2d31807cbf6 clarified type ML_Context.antiq: context parser maintains compilation context, declaration is applied to final context;
wenzelm
parents: 53168
diff changeset
    60
    val initial = null (get_thmss ctxt);
59112
e670969f34df expand ML cartouches to Input.source;
wenzelm
parents: 58866
diff changeset
    61
    val (name, ctxt') = ML_Context.variant kind ctxt;
53169
e2d31807cbf6 clarified type ML_Context.antiq: context parser maintains compilation context, declaration is applied to final context;
wenzelm
parents: 53168
diff changeset
    62
    val ctxt'' = cons_thms ((name, is_single), thms) ctxt';
27389
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    63
59127
723b11f8ffbf more careful handling of auxiliary environment structure -- allow nested ML evaluation;
wenzelm
parents: 59112
diff changeset
    64
    val ml_body = ML_Context.struct_name ctxt ^ "." ^ name;
53169
e2d31807cbf6 clarified type ML_Context.antiq: context parser maintains compilation context, declaration is applied to final context;
wenzelm
parents: 53168
diff changeset
    65
    fun decl final_ctxt =
48782
e955964d89cb more direct embedding of abstract thm values into the ML environment -- avoidance of repeated ML_Thms.the_thm(s) considerably reduces compilation time for Poly/ML 5.4.x;
wenzelm
parents: 48776
diff changeset
    66
      if initial then
e955964d89cb more direct embedding of abstract thm values into the ML environment -- avoidance of repeated ML_Thms.the_thm(s) considerably reduces compilation time for Poly/ML 5.4.x;
wenzelm
parents: 48776
diff changeset
    67
        let
53169
e2d31807cbf6 clarified type ML_Context.antiq: context parser maintains compilation context, declaration is applied to final context;
wenzelm
parents: 53168
diff changeset
    68
          val binds = get_thmss final_ctxt |> map (fn ((a, b), _) => (b ? enclose "[" "]") a);
48782
e955964d89cb more direct embedding of abstract thm values into the ML environment -- avoidance of repeated ML_Thms.the_thm(s) considerably reduces compilation time for Poly/ML 5.4.x;
wenzelm
parents: 48776
diff changeset
    69
          val ml_env = "val [" ^ commas binds ^ "] = ML_Thms.the_thmss ML_context;\n";
e955964d89cb more direct embedding of abstract thm values into the ML environment -- avoidance of repeated ML_Thms.the_thm(s) considerably reduces compilation time for Poly/ML 5.4.x;
wenzelm
parents: 48776
diff changeset
    70
        in (ml_env, ml_body) end
e955964d89cb more direct embedding of abstract thm values into the ML environment -- avoidance of repeated ML_Thms.the_thm(s) considerably reduces compilation time for Poly/ML 5.4.x;
wenzelm
parents: 48776
diff changeset
    71
      else ("", ml_body);
56069
451d5b73f8cf simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents: 55997
diff changeset
    72
  in (decl, ctxt'') end;
27389
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    73
53171
a5e54d4d9081 added Theory.setup convenience;
wenzelm
parents: 53169
diff changeset
    74
val _ = Theory.setup
67147
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 66067
diff changeset
    75
  (ML_Antiquotation.declaration \<^binding>\<open>thm\<close> (Attrib.thm >> single) (K (thm_binding "thm" true)) #>
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 66067
diff changeset
    76
   ML_Antiquotation.declaration \<^binding>\<open>thms\<close> Attrib.thms (K (thm_binding "thms" false)));
27389
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    77
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    78
74603
c22ae7b41bb8 clarified signature;
wenzelm
parents: 74602
diff changeset
    79
(* embedded lemma *)
c22ae7b41bb8 clarified signature;
wenzelm
parents: 74602
diff changeset
    80
c22ae7b41bb8 clarified signature;
wenzelm
parents: 74602
diff changeset
    81
val embedded_lemma =
74604
3da2662a35cd local fixes for "lemma" antiquotation;
wenzelm
parents: 74603
diff changeset
    82
  Args.mode "open" -- Parse.and_list1 (Scan.repeat1 Parse.embedded_inner_syntax)
3da2662a35cd local fixes for "lemma" antiquotation;
wenzelm
parents: 74603
diff changeset
    83
    -- Parse.for_fixes -- Method.parse_by
3da2662a35cd local fixes for "lemma" antiquotation;
wenzelm
parents: 74603
diff changeset
    84
    >> (fn (((is_open, raw_stmt), fixes), (methods, reports)) => fn ctxt =>
74603
c22ae7b41bb8 clarified signature;
wenzelm
parents: 74602
diff changeset
    85
        let
c22ae7b41bb8 clarified signature;
wenzelm
parents: 74602
diff changeset
    86
          val _ = Context_Position.reports ctxt reports;
c22ae7b41bb8 clarified signature;
wenzelm
parents: 74602
diff changeset
    87
74606
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74604
diff changeset
    88
          val fixes_ctxt = #2 (Proof_Context.add_fixes_cmd fixes ctxt);
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74604
diff changeset
    89
          val stmt = burrow (map (rpair []) o Syntax.read_props fixes_ctxt) raw_stmt;
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74604
diff changeset
    90
          val stmt_ctxt = (fold o fold) (Proof_Context.augment o #1) stmt fixes_ctxt;
74604
3da2662a35cd local fixes for "lemma" antiquotation;
wenzelm
parents: 74603
diff changeset
    91
74603
c22ae7b41bb8 clarified signature;
wenzelm
parents: 74602
diff changeset
    92
          val prep_result = Goal.norm_result ctxt #> not is_open ? Thm.close_derivation \<^here>;
c22ae7b41bb8 clarified signature;
wenzelm
parents: 74602
diff changeset
    93
          fun after_qed res goal_ctxt =
c22ae7b41bb8 clarified signature;
wenzelm
parents: 74602
diff changeset
    94
            Proof_Context.put_thms false (Auto_Bind.thisN,
c22ae7b41bb8 clarified signature;
wenzelm
parents: 74602
diff changeset
    95
              SOME (map prep_result (Proof_Context.export goal_ctxt ctxt (flat res)))) goal_ctxt;
c22ae7b41bb8 clarified signature;
wenzelm
parents: 74602
diff changeset
    96
74604
3da2662a35cd local fixes for "lemma" antiquotation;
wenzelm
parents: 74603
diff changeset
    97
          val thms_ctxt = stmt_ctxt
74603
c22ae7b41bb8 clarified signature;
wenzelm
parents: 74602
diff changeset
    98
            |> Proof.theorem NONE after_qed stmt
c22ae7b41bb8 clarified signature;
wenzelm
parents: 74602
diff changeset
    99
            |> Proof.global_terminal_proof methods;
c22ae7b41bb8 clarified signature;
wenzelm
parents: 74602
diff changeset
   100
          val thms =
74604
3da2662a35cd local fixes for "lemma" antiquotation;
wenzelm
parents: 74603
diff changeset
   101
            Proof_Context.get_fact thms_ctxt
3da2662a35cd local fixes for "lemma" antiquotation;
wenzelm
parents: 74603
diff changeset
   102
              (Facts.named (Proof_Context.full_name thms_ctxt (Binding.name Auto_Bind.thisN)))
74606
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74604
diff changeset
   103
        in (thms, (map #1 (flat stmt), stmt_ctxt)) end);
27389
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
   104
53171
a5e54d4d9081 added Theory.setup convenience;
wenzelm
parents: 53169
diff changeset
   105
val _ = Theory.setup
74603
c22ae7b41bb8 clarified signature;
wenzelm
parents: 74602
diff changeset
   106
  (ML_Antiquotation.declaration \<^binding>\<open>lemma\<close> (Scan.lift embedded_lemma)
c22ae7b41bb8 clarified signature;
wenzelm
parents: 74602
diff changeset
   107
    (fn _ => fn make_lemma => fn ctxt =>
74606
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74604
diff changeset
   108
      let val thms = #1 (make_lemma ctxt)
74596
55d4f8e1877f tuned, continuing e955964d89cb;
wenzelm
parents: 74563
diff changeset
   109
      in thm_binding "lemma" (length thms = 1) thms ctxt end));
27389
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
   110
56199
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 56072
diff changeset
   111
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 56072
diff changeset
   112
(* old-style theorem bindings *)
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 56072
diff changeset
   113
72054
6c75287276d5 tuned -- avoid non-standard extend/merge;
wenzelm
parents: 70494
diff changeset
   114
structure Data = Theory_Data
56199
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 56072
diff changeset
   115
(
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 56072
diff changeset
   116
  type T = thm list;
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 56072
diff changeset
   117
  val empty = [];
72054
6c75287276d5 tuned -- avoid non-standard extend/merge;
wenzelm
parents: 70494
diff changeset
   118
  val merge = #1;
56199
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 56072
diff changeset
   119
);
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 56072
diff changeset
   120
72054
6c75287276d5 tuned -- avoid non-standard extend/merge;
wenzelm
parents: 70494
diff changeset
   121
fun get_stored_thms () = Data.get (Context.the_global_context ());
56199
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 56072
diff changeset
   122
val get_stored_thm = hd o get_stored_thms;
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 56072
diff changeset
   123
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 56072
diff changeset
   124
fun ml_store get (name, ths) =
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 56072
diff changeset
   125
  let
67715
ec46ecb87999 tuned signature;
wenzelm
parents: 67147
diff changeset
   126
    val (_, ths') =
78795
f7e972d567f3 clarified signature: more concise variations on implicit theory setup;
wenzelm
parents: 74606
diff changeset
   127
      Theory.setup_result (Global_Theory.note_thms "" ((Binding.name name, []), [(ths, [])]));
72054
6c75287276d5 tuned -- avoid non-standard extend/merge;
wenzelm
parents: 70494
diff changeset
   128
    val _ = Theory.setup (Data.put ths');
56199
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 56072
diff changeset
   129
    val _ =
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 56072
diff changeset
   130
      if name = "" then ()
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 56072
diff changeset
   131
      else if not (ML_Syntax.is_identifier name) then
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 56072
diff changeset
   132
        error ("Cannot bind theorem(s) " ^ quote name ^ " as ML value")
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 56072
diff changeset
   133
      else
56304
40274e4f5ebf redirect ML_Compiler reports more directly: only the (big) parse tree report is deferred via Execution.print (NB: this does not work for asynchronous "diag" commands);
wenzelm
parents: 56275
diff changeset
   134
        ML_Compiler.eval (ML_Compiler.verbose true ML_Compiler.flags) Position.none
56199
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 56072
diff changeset
   135
          (ML_Lex.tokenize ("val " ^ name ^ " = " ^ get ^ " ();"));
72054
6c75287276d5 tuned -- avoid non-standard extend/merge;
wenzelm
parents: 70494
diff changeset
   136
    val _ = Theory.setup (Data.put []);
56199
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 56072
diff changeset
   137
  in () end;
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 56072
diff changeset
   138
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 56072
diff changeset
   139
val store_thms = ml_store "ML_Thms.get_stored_thms";
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 56072
diff changeset
   140
fun store_thm (name, th) = ml_store "ML_Thms.get_stored_thm" (name, [th]);
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 56072
diff changeset
   141
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 56072
diff changeset
   142
fun bind_thm (name, thm) = store_thm (name, Drule.export_without_context thm);
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 56072
diff changeset
   143
fun bind_thms (name, thms) = store_thms (name, map Drule.export_without_context thms);
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 56072
diff changeset
   144
27389
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
   145
end;