src/Pure/ML/ml_thms.ML
author wenzelm
Sat, 13 Dec 2008 15:00:39 +0100
changeset 29091 b81fe045e799
parent 27869 af1f79cbc198
child 29606 fedb8be05f24
permissions -rw-r--r--
Context.display_names;
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
    ID:         $Id$
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
     3
    Author:     Makarius
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
     4
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
     5
Isar theorem values within ML.
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
     6
*)
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
     7
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
     8
signature ML_THMS =
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
     9
sig
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    10
  val the_thms: Proof.context -> int -> thm list
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    11
  val the_thm: Proof.context -> int -> thm
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    12
end;
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    13
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    14
structure ML_Thms: ML_THMS =
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    15
struct
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    16
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    17
(* auxiliary facts table *)
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    18
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    19
structure AuxFacts = ProofDataFun
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    20
(
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    21
  type T = thm list Inttab.table;
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    22
  fun init _ = Inttab.empty;
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    23
);
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    24
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    25
val put_thms = AuxFacts.map o Inttab.update;
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    26
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    27
fun the_thms ctxt name = the (Inttab.lookup (AuxFacts.get ctxt) name);
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    28
fun the_thm ctxt name = the_single (the_thms ctxt name);
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    29
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    30
fun thm_bind kind a i =
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    31
  "val " ^ a ^ " = ML_Thms.the_" ^ kind ^
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    32
    " (ML_Context.the_local_context ()) " ^ string_of_int i ^ ";\n";
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    33
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    34
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    35
(* fact references *)
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    36
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    37
fun thm_antiq kind scan = ML_Context.add_antiq kind
27869
af1f79cbc198 ML_Context.add_antiq: pass position;
wenzelm
parents: 27809
diff changeset
    38
  (fn _ => scan >> (fn ths => fn {struct_name, background} =>
27389
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    39
    let
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    40
      val i = serial ();
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    41
      val (a, background') = background
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    42
        |> ML_Antiquote.variant kind ||> put_thms (i, ths);
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    43
      val ml = (thm_bind kind a i, struct_name ^ "." ^ a);
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    44
    in (K ml, background') end));
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    45
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    46
val _ = thm_antiq "thm" (Attrib.thm >> single);
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    47
val _ = thm_antiq "thms" Attrib.thms;
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    48
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    49
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    50
(* ad-hoc goals *)
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    51
27809
a1e409db516b unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents: 27521
diff changeset
    52
val by = Args.$$$ "by";
a1e409db516b unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents: 27521
diff changeset
    53
val goal = Scan.unless (Scan.lift by) Args.prop;
27389
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    54
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    55
val _ = ML_Context.add_antiq "lemma"
27869
af1f79cbc198 ML_Context.add_antiq: pass position;
wenzelm
parents: 27809
diff changeset
    56
  (fn pos => Args.context -- Args.mode "open" -- Scan.repeat1 goal --
27809
a1e409db516b unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents: 27521
diff changeset
    57
      Scan.lift (by |-- Method.parse -- Scan.option Method.parse) >>
27521
44081396d735 @{lemma}: allow terminal method, close derivation unless (open) mode is given;
wenzelm
parents: 27389
diff changeset
    58
    (fn (((ctxt, is_open), props), methods) => fn {struct_name, background} =>
27389
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    59
      let
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    60
        val i = serial ();
27869
af1f79cbc198 ML_Context.add_antiq: pass position;
wenzelm
parents: 27809
diff changeset
    61
        val prep_result =
af1f79cbc198 ML_Context.add_antiq: pass position;
wenzelm
parents: 27809
diff changeset
    62
          Goal.norm_result #> Thm.default_position pos #> not is_open ? Thm.close_derivation;
27389
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    63
        fun after_qed [res] goal_ctxt =
27521
44081396d735 @{lemma}: allow terminal method, close derivation unless (open) mode is given;
wenzelm
parents: 27389
diff changeset
    64
          put_thms (i, map prep_result (ProofContext.export goal_ctxt ctxt res)) goal_ctxt;
27389
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    65
        val ctxt' = ctxt
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    66
          |> Proof.theorem_i NONE after_qed [map (rpair []) props]
27521
44081396d735 @{lemma}: allow terminal method, close derivation unless (open) mode is given;
wenzelm
parents: 27389
diff changeset
    67
          |> Proof.global_terminal_proof methods;
27389
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    68
        val (a, background') = background
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    69
          |> ML_Antiquote.variant "lemma" ||> put_thms (i, the_thms ctxt' i);
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    70
        val ml = (thm_bind (if length props = 1 then "thm" else "thms") a i, struct_name ^ "." ^ a);
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    71
      in (K ml, background') end));
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    72
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    73
end;
823c7ec3ea4f Isar theorem values within ML.
wenzelm
parents:
diff changeset
    74