src/Pure/ML/ml_instantiate.ML
author wenzelm
Fri, 29 Oct 2021 11:59:02 +0200
changeset 74620 d622d1dce05c
parent 74606 40f5c6b2e8aa
child 78804 d4e9d6b7f48d
permissions -rw-r--r--
avoid persistence of static context: instantiation arguments should provide proper dynamic context (see also e2e2bc1f9570);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
74593
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/ML/ml_instantiate.ML
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
     3
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
     4
ML antiquotation to instantiate logical entities.
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
     5
*)
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
     6
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
     7
signature ML_INSTANTIATE =
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
     8
sig
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
     9
  val make_ctyp: Proof.context -> typ -> ctyp
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    10
  val make_cterm: Proof.context -> term -> cterm
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    11
  type insts = ((indexname * sort) * typ) list * ((indexname * typ) * term) list
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    12
  type cinsts = ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    13
  val instantiate_typ: insts -> typ -> typ
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    14
  val instantiate_ctyp: Position.T -> cinsts -> ctyp -> ctyp
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    15
  val instantiate_term: insts -> term -> term
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    16
  val instantiate_cterm: Position.T -> cinsts -> cterm -> cterm
74606
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
    17
  val instantiate_thm: Position.T -> cinsts -> thm -> thm
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
    18
  val instantiate_thms: Position.T -> cinsts -> thm list -> thm list
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
    19
  val get_thms: Proof.context -> int -> thm list
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
    20
  val get_thm: Proof.context -> int -> thm
74593
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    21
end;
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    22
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    23
structure ML_Instantiate: ML_INSTANTIATE =
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    24
struct
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    25
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    26
(* exported operations *)
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    27
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    28
fun make_ctyp ctxt T = Thm.ctyp_of ctxt T |> Thm.trim_context_ctyp;
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    29
fun make_cterm ctxt t = Thm.cterm_of ctxt t |> Thm.trim_context_cterm;
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    30
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    31
type insts = ((indexname * sort) * typ) list * ((indexname * typ) * term) list
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    32
type cinsts = ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    33
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    34
fun instantiate_typ (insts: insts) =
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    35
  Term_Subst.instantiateT (TVars.make (#1 insts));
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    36
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    37
fun instantiate_ctyp pos (cinsts: cinsts) cT =
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    38
  Thm.instantiate_ctyp (TVars.make (#1 cinsts)) cT
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    39
  handle CTERM (msg, args) => Exn.reraise (CTERM (msg ^ Position.here pos, args));
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    40
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    41
fun instantiate_term (insts: insts) =
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    42
  let
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    43
    val instT = TVars.make (#1 insts);
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    44
    val instantiateT = Term_Subst.instantiateT instT;
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    45
    val inst = Vars.make ((map o apfst o apsnd) instantiateT (#2 insts));
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    46
  in Term_Subst.instantiate_beta (instT, inst) end;
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    47
74606
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
    48
fun make_cinsts (cinsts: cinsts) =
74593
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    49
  let
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    50
    val cinstT = TVars.make (#1 cinsts);
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    51
    val instantiateT = Term_Subst.instantiateT (TVars.map (K Thm.typ_of) cinstT);
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    52
    val cinst = Vars.make ((map o apfst o apsnd) instantiateT (#2 cinsts));
74606
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
    53
  in (cinstT, cinst) end;
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
    54
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
    55
fun instantiate_cterm pos cinsts ct =
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
    56
  Thm.instantiate_beta_cterm (make_cinsts cinsts) ct
74593
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    57
  handle CTERM (msg, args) => Exn.reraise (CTERM (msg ^ Position.here pos, args));
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    58
74606
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
    59
fun instantiate_thm pos cinsts th =
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
    60
  Thm.instantiate_beta (make_cinsts cinsts) th
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
    61
  handle THM (msg, i, args) => Exn.reraise (THM (msg ^ Position.here pos, i, args));
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
    62
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
    63
fun instantiate_thms pos cinsts = map (instantiate_thm pos cinsts);
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
    64
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
    65
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
    66
(* context data *)
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
    67
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
    68
structure Data = Proof_Data
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
    69
(
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
    70
  type T = int * thm list Inttab.table;
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
    71
  fun init _ = (0, Inttab.empty);
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
    72
);
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
    73
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
    74
fun put_thms ths ctxt =
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
    75
  let
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
    76
    val (i, thms) = Data.get ctxt;
74620
d622d1dce05c avoid persistence of static context: instantiation arguments should provide proper dynamic context (see also e2e2bc1f9570);
wenzelm
parents: 74606
diff changeset
    77
    val ctxt' = ctxt |> Data.put (i + 1, Inttab.update (i, map Thm.trim_context ths) thms);
74606
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
    78
  in (i, ctxt') end;
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
    79
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
    80
fun get_thms ctxt i = the (Inttab.lookup (#2 (Data.get ctxt)) i);
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
    81
fun get_thm ctxt i = the_single (get_thms ctxt i);
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
    82
74593
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    83
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    84
(* ML antiquotation *)
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    85
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    86
local
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    87
74605
wenzelm
parents: 74593
diff changeset
    88
val make_keywords =
wenzelm
parents: 74593
diff changeset
    89
  Thy_Header.get_keywords'
wenzelm
parents: 74593
diff changeset
    90
  #> Keyword.no_major_keywords
74606
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
    91
  #> Keyword.add_major_keywords ["typ", "term", "prop", "ctyp", "cterm", "cprop", "lemma"];
74593
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    92
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    93
val parse_inst_name = Parse.position (Parse.type_ident >> pair true || Parse.name >> pair false);
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    94
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    95
val parse_inst =
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    96
  (parse_inst_name -- (Parse.$$$ "=" |-- Parse.!!! Parse.embedded_ml) ||
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    97
    Scan.ahead parse_inst_name -- Parse.embedded_ml)
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    98
  >> (fn (((b, a), pos), ml) => (b, ((a, pos), ml)));
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
    99
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   100
val parse_insts =
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   101
  Parse.and_list1 parse_inst >> (List.partition #1 #> apply2 (map #2));
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   102
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   103
val ml = ML_Lex.tokenize_no_range;
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   104
val ml_range = ML_Lex.tokenize_range;
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   105
fun ml_parens x = ml "(" @ x @ ml ")";
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   106
fun ml_bracks x = ml "[" @ x @ ml "]";
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   107
fun ml_commas xs = flat (separate (ml ", ") xs);
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   108
val ml_list = ml_bracks o ml_commas;
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   109
fun ml_pair (x, y) = ml_parens (ml_commas [x, y]);
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   110
val ml_list_pair = ml_list o ListPair.map ml_pair;
74606
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   111
val ml_here = ML_Syntax.atomic o ML_Syntax.print_position;
74593
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   112
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   113
fun get_tfree envT (a, pos) =
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   114
  (case AList.lookup (op =) envT a of
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   115
    SOME S => (a, S)
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   116
  | NONE => error ("No occurrence of type variable " ^ quote a ^ Position.here pos));
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   117
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   118
fun check_free ctxt env (x, pos) =
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   119
  (case AList.lookup (op =) env x of
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   120
    SOME T =>
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   121
      (Context_Position.reports ctxt (map (pair pos) (Syntax_Phases.markup_free ctxt x)); (x, T))
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   122
  | NONE => error ("No occurrence of variable " ^ quote x ^ Position.here pos));
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   123
74606
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   124
fun missing_instT pos envT instT =
74593
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   125
  (case filter_out (fn (a, _) => exists (fn (b, _) => a = b) instT) envT of
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   126
    [] => ()
74606
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   127
  | bad =>
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   128
      error ("No instantiation for free type variable(s) " ^ commas_quote (map #1 bad) ^
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   129
        Position.here pos));
74593
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   130
74606
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   131
fun missing_inst pos env inst =
74593
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   132
  (case filter_out (fn (a, _) => exists (fn (b, _) => a = b) inst) env of
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   133
    [] => ()
74606
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   134
  | bad =>
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   135
      error ("No instantiation for free variable(s) " ^ commas_quote (map #1 bad) ^
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   136
        Position.here pos));
74593
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   137
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   138
fun make_instT (a, pos) T =
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   139
  (case try (Term.dest_TVar o Logic.dest_type) T of
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   140
    NONE => error ("Not a free type variable " ^ quote a ^ Position.here pos)
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   141
  | SOME v => ml (ML_Syntax.print_pair ML_Syntax.print_indexname ML_Syntax.print_sort v));
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   142
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   143
fun make_inst (a, pos) t =
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   144
  (case try Term.dest_Var t of
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   145
    NONE => error ("Not a free variable " ^ quote a ^ Position.here pos)
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   146
  | SOME v => ml (ML_Syntax.print_pair ML_Syntax.print_indexname ML_Syntax.print_typ v));
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   147
74606
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   148
fun make_env ts =
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   149
  let
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   150
    val envT = fold Term.add_tfrees ts [];
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   151
    val env = fold Term.add_frees ts [];
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   152
  in (envT, env) end;
74593
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   153
74606
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   154
fun prepare_insts pos {schematic} ctxt1 ctxt0 (instT, inst) ts =
74593
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   155
  let
74606
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   156
    val (envT, env) = make_env ts;
74593
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   157
    val freesT = map (Logic.mk_type o TFree o get_tfree envT) instT;
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   158
    val frees = map (Free o check_free ctxt1 env) inst;
74606
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   159
    val (ts', (varsT, vars)) =
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   160
      Variable.export_terms ctxt1 ctxt0 (ts @ freesT @ frees)
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   161
      |> chop (length ts) ||> chop (length freesT);
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   162
    val ml_insts = (map2 make_instT instT varsT, map2 make_inst inst vars);
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   163
  in
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   164
    if schematic then ()
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   165
    else
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   166
      let val (envT', env') = make_env ts' in
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   167
        missing_instT pos (subtract (eq_fst op =) envT' envT) instT;
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   168
        missing_inst pos (subtract (eq_fst op =) env' env) inst
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   169
      end;
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   170
    (ml_insts, ts')
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   171
  end;
74593
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   172
74606
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   173
fun prepare_ml range kind ml1 ml2 ml_val (ml_instT, ml_inst) ctxt =
74593
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   174
  let
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   175
    val (ml_name, ctxt') = ML_Context.variant kind ctxt;
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   176
    val ml_env = ml ("val " ^ ml_name ^ " = ") @ ml ml1 @ ml_parens (ml ml_val) @ ml ";\n";
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   177
    fun ml_body (ml_argsT, ml_args) =
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   178
      ml_parens (ml ml2 @
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   179
        ml_pair (ml_list_pair (ml_instT, ml_argsT), ml_list_pair (ml_inst, ml_args)) @
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   180
        ml_range range (ML_Context.struct_name ctxt ^ "." ^ ml_name));
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   181
  in ((ml_env, ml_body), ctxt') end;
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   182
74606
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   183
fun prepare_type range ((((kind, pos), ml1, ml2), schematic), s) insts ctxt =
74593
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   184
  let
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   185
    val T = Syntax.read_typ ctxt s;
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   186
    val t = Logic.mk_type T;
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   187
    val ctxt1 = Proof_Context.augment t ctxt;
74606
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   188
    val (ml_insts, T') =
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   189
      prepare_insts pos schematic ctxt1 ctxt insts [t] ||> (the_single #> Logic.dest_type);
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   190
  in prepare_ml range kind ml1 ml2 (ML_Syntax.print_typ T') ml_insts ctxt end;
74593
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   191
74606
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   192
fun prepare_term read range ((((kind, pos), ml1, ml2), schematic), (s, fixes)) insts ctxt =
74593
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   193
  let
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   194
    val ctxt' = #2 (Proof_Context.add_fixes_cmd fixes ctxt);
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   195
    val t = read ctxt' s;
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   196
    val ctxt1 = Proof_Context.augment t ctxt';
74606
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   197
    val (ml_insts, t') = prepare_insts pos schematic ctxt1 ctxt insts [t] ||> the_single;
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   198
  in prepare_ml range kind ml1 ml2 (ML_Syntax.print_term t') ml_insts ctxt end;
74593
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   199
74606
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   200
fun prepare_lemma range ((pos, schematic), make_lemma) insts ctxt =
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   201
  let
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   202
    val (ths, (props, ctxt1)) = make_lemma ctxt
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   203
    val (i, thms_ctxt) = put_thms ths ctxt;
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   204
    val ml_insts = #1 (prepare_insts pos schematic ctxt1 ctxt insts props);
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   205
    val (ml1, ml2) =
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   206
      if length ths = 1
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   207
      then ("ML_Instantiate.get_thm ML_context", "ML_Instantiate.instantiate_thm " ^ ml_here pos)
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   208
      else ("ML_Instantiate.get_thms ML_context", "ML_Instantiate.instantiate_thms " ^ ml_here pos);
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   209
  in prepare_ml range "lemma" ml1 ml2 (ML_Syntax.print_int i) ml_insts thms_ctxt end;
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   210
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   211
fun typ_ml (kind, pos: Position.T) = ((kind, pos), "", "ML_Instantiate.instantiate_typ ");
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   212
fun term_ml (kind, pos: Position.T) = ((kind, pos), "", "ML_Instantiate.instantiate_term ");
74593
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   213
fun ctyp_ml (kind, pos) =
74606
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   214
  ((kind, pos),
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   215
    "ML_Instantiate.make_ctyp ML_context", "ML_Instantiate.instantiate_ctyp " ^ ml_here pos);
74593
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   216
fun cterm_ml (kind, pos) =
74606
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   217
  ((kind, pos),
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   218
    "ML_Instantiate.make_cterm ML_context", "ML_Instantiate.instantiate_cterm " ^ ml_here pos);
74593
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   219
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   220
val command_name = Parse.position o Parse.command_name;
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   221
74606
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   222
val parse_schematic = Args.mode "schematic" >> (fn b => {schematic = b});
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   223
74593
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   224
fun parse_body range =
74606
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   225
  (command_name "typ" >> typ_ml || command_name "ctyp" >> ctyp_ml) -- parse_schematic --
74593
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   226
    Parse.!!! Parse.typ >> prepare_type range ||
74606
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   227
  (command_name "term" >> term_ml || command_name "cterm" >> cterm_ml) -- parse_schematic --
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   228
    Parse.!!! (Parse.term -- Parse.for_fixes) >> prepare_term Syntax.read_term range ||
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   229
  (command_name "prop" >> term_ml || command_name "cprop" >> cterm_ml) -- parse_schematic --
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   230
    Parse.!!! (Parse.term -- Parse.for_fixes) >> prepare_term Syntax.read_prop range ||
40f5c6b2e8aa support for "lemma";
wenzelm
parents: 74605
diff changeset
   231
  (command_name "lemma" >> #2) -- parse_schematic -- ML_Thms.embedded_lemma >> prepare_lemma range;
74593
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   232
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   233
val _ = Theory.setup
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   234
  (ML_Context.add_antiquotation \<^binding>\<open>instantiate\<close> true
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   235
    (fn range => fn src => fn ctxt =>
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   236
      let
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   237
        val (insts, prepare_val) = src
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   238
          |> Parse.read_embedded_src ctxt (make_keywords ctxt)
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   239
              ((parse_insts --| Parse.$$$ "in") -- parse_body range);
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   240
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   241
        val (((ml_env, ml_body), decls), ctxt1) = ctxt
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   242
          |> prepare_val (apply2 (map #1) insts)
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   243
          ||>> ML_Context.expand_antiquotes_list (op @ (apply2 (map #2) insts));
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   244
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   245
        fun decl' ctxt' =
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   246
          let val (ml_args_env, ml_args) = split_list (decls ctxt')
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   247
          in (ml_env @ flat ml_args_env, ml_body (chop (length (#1 insts)) ml_args)) end;
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   248
      in (decl', ctxt1) end));
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   249
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   250
in end;
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   251
66f10c877542 clarified modules;
wenzelm
parents:
diff changeset
   252
end;