src/Pure/Tools/invoke.ML
author wenzelm
Tue, 10 Jul 2007 23:29:43 +0200
changeset 23719 ccd9cb15c062
parent 23351 3702086a15a3
child 24493 d4380e9b287b
permissions -rw-r--r--
more markup for inner and outer syntax; added enclose;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
19813
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Tools/invoke.ML
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
     3
    Author:     Makarius
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
     4
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
     5
Schematic invocation of locale expression in proof context.
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
     6
*)
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
     7
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
     8
signature INVOKE =
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
     9
sig
19849
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    10
  val invoke: string * Attrib.src list -> Locale.expr -> string option list ->
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    11
    (string * string option * mixfix) list -> bool -> Proof.state -> Proof.state
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    12
  val invoke_i: string * attribute list -> Locale.expr -> term option list ->
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    13
    (string * typ option * mixfix) list -> bool -> Proof.state -> Proof.state
19813
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
    14
end;
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
    15
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
    16
structure Invoke: INVOKE =
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
    17
struct
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
    18
19849
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    19
19813
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
    20
(* invoke *)
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
    21
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
    22
local
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
    23
19849
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    24
fun gen_invoke prep_att prep_expr prep_terms add_fixes
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    25
    (prfx, raw_atts) raw_expr raw_insts fixes int state =
19813
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
    26
  let
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
    27
    val _ = Proof.assert_forward_or_chain state;
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
    28
    val thy = Proof.theory_of state;
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
    29
19849
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    30
    val more_atts = map (prep_att thy) raw_atts;
19813
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
    31
    val (elems, _) = prep_expr raw_expr [] (ProofContext.init thy);
19849
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    32
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    33
    val prems = maps Element.prems_of elems;
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    34
    val params = maps Element.params_of elems;
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    35
    val types = rev (fold Term.add_tfrees prems (fold (Term.add_tfreesT o #2) params []));
19813
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
    36
19849
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    37
    val prems' = map Logic.varify prems;
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    38
    val params' = map (Logic.varify o Free) params;
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    39
    val types' = map (Logic.varifyT o TFree) types;
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    40
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    41
    val state' = state
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    42
      |> Proof.begin_block
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    43
      |> Proof.map_context (snd o add_fixes fixes);
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    44
    val ctxt' = Proof.context_of state';
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    45
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    46
    val raw_insts' = zip_options params' raw_insts
19813
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
    47
      handle Library.UnequalLengths => error "Too many instantiations";
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
    48
    val insts = map #1 raw_insts' ~~
22769
6f5068e26b89 simplified ProofContext.read_termTs;
wenzelm
parents: 22658
diff changeset
    49
      Variable.polymorphic ctxt'
6f5068e26b89 simplified ProofContext.read_termTs;
wenzelm
parents: 22658
diff changeset
    50
        (prep_terms ctxt' (map (fn (t, u) => (u, Term.fastype_of t)) raw_insts'));
6f5068e26b89 simplified ProofContext.read_termTs;
wenzelm
parents: 22658
diff changeset
    51
19849
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    52
    val inst_rules =
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    53
      replicate (length types') Drule.termI @
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    54
      map (fn t =>
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    55
        (case AList.lookup (op =) insts t of
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    56
          SOME u => Drule.mk_term (Thm.cterm_of thy u)
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    57
        | NONE => Drule.termI)) params';
19813
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
    58
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
    59
    val propp =
19849
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    60
      [(("", []), map (rpair [] o Logic.mk_term o Logic.mk_type) types'),
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    61
       (("", []), map (rpair [] o Logic.mk_term) params'),
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    62
       (("", []), map (rpair [] o Element.mark_witness) prems')];
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    63
    fun after_qed results =
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    64
      Proof.end_block #>
20305
16c8f44b1852 simplified Proof.end_block;
wenzelm
parents: 20266
diff changeset
    65
      Proof.map_context (fn ctxt =>
19849
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    66
        let
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    67
          val ([res_types, res_params, res_prems], ctxt'') =
22568
ed7aa5a350ef renamed Variable.import to import_thms (avoid clash with Alice keywords);
wenzelm
parents: 22101
diff changeset
    68
            fold_burrow (apfst snd oo Variable.import_thms false) results ctxt';
19849
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    69
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    70
          val types'' = map (Logic.dest_type o Thm.term_of o Drule.dest_term) res_types;
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    71
          val params'' = map (Thm.term_of o Drule.dest_term) res_params;
21483
e4be91feca50 replaced Args.map_values/Element.map_ctxt_values by general morphism application;
wenzelm
parents: 21440
diff changeset
    72
          val inst = Element.morph_ctxt (Element.inst_morphism thy
19849
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    73
            (Symtab.make (map #1 types ~~ types''), Symtab.make (map #1 params ~~ params'')));
21483
e4be91feca50 replaced Args.map_values/Element.map_ctxt_values by general morphism application;
wenzelm
parents: 21440
diff changeset
    74
          val elems' = map inst elems;
19868
wenzelm
parents: 19852
diff changeset
    75
          val prems'' = map2 Element.make_witness (maps Element.prems_of elems') res_prems;
20266
6fb734789818 proper Element.generalize_facts;
wenzelm
parents: 20218
diff changeset
    76
          val notes =
6fb734789818 proper Element.generalize_facts;
wenzelm
parents: 20218
diff changeset
    77
            maps (Element.facts_of thy) elems'
6fb734789818 proper Element.generalize_facts;
wenzelm
parents: 20218
diff changeset
    78
            |> Element.satisfy_facts prems''
21584
22c9392de970 Element.generalize_facts;
wenzelm
parents: 21483
diff changeset
    79
            |> Element.generalize_facts ctxt'' ctxt
19849
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    80
            |> Attrib.map_facts (Attrib.attribute_i thy)
20266
6fb734789818 proper Element.generalize_facts;
wenzelm
parents: 20218
diff changeset
    81
            |> map (fn ((a, atts), bs) => ((a, atts @ more_atts), bs));
19849
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    82
        in
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    83
          ctxt
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    84
          |> ProofContext.sticky_prefix prfx
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    85
          |> ProofContext.qualified_names
21440
807a39221a58 notes: proper kind;
wenzelm
parents: 20893
diff changeset
    86
          |> (snd o ProofContext.note_thmss_i "" notes)
19849
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    87
          |> ProofContext.restore_naming ctxt
20305
16c8f44b1852 simplified Proof.end_block;
wenzelm
parents: 20266
diff changeset
    88
        end) #>
16c8f44b1852 simplified Proof.end_block;
wenzelm
parents: 20266
diff changeset
    89
      Proof.put_facts NONE #> Seq.single;
19813
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
    90
  in
19849
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
    91
    state'
19868
wenzelm
parents: 19852
diff changeset
    92
    |> Proof.local_goal (K (K ())) (K I) ProofContext.bind_propp_schematic_i
19813
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
    93
      "invoke" NONE after_qed propp
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
    94
    |> Element.refine_witness
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
    95
    |> Seq.hd
23351
3702086a15a3 Method.Basic: include position;
wenzelm
parents: 22769
diff changeset
    96
    |> Proof.refine (Method.Basic (K (Method.METHOD (K (HEADGOAL (RANGE (map rtac inst_rules))))),
3702086a15a3 Method.Basic: include position;
wenzelm
parents: 22769
diff changeset
    97
      Position.none))
19813
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
    98
    |> Seq.hd
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
    99
  end;
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
   100
22769
6f5068e26b89 simplified ProofContext.read_termTs;
wenzelm
parents: 22658
diff changeset
   101
fun infer_terms ctxt =
6f5068e26b89 simplified ProofContext.read_termTs;
wenzelm
parents: 22658
diff changeset
   102
  ProofContext.infer_types ctxt o
6f5068e26b89 simplified ProofContext.read_termTs;
wenzelm
parents: 22658
diff changeset
   103
    (map (fn (t, T) => TypeInfer.constrain t (TypeInfer.paramify_vars T)));
19813
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
   104
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
   105
in
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
   106
22769
6f5068e26b89 simplified ProofContext.read_termTs;
wenzelm
parents: 22658
diff changeset
   107
fun invoke x =
6f5068e26b89 simplified ProofContext.read_termTs;
wenzelm
parents: 22658
diff changeset
   108
  gen_invoke Attrib.attribute Locale.read_expr ProofContext.read_termTs ProofContext.add_fixes x;
6f5068e26b89 simplified ProofContext.read_termTs;
wenzelm
parents: 22658
diff changeset
   109
fun invoke_i x = gen_invoke (K I) Locale.cert_expr infer_terms ProofContext.add_fixes_i x;
19813
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
   110
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
   111
end;
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
   112
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
   113
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
   114
(* concrete syntax *)
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
   115
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
   116
local structure P = OuterParse and K = OuterKeyword in
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
   117
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
   118
val invokeP =
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
   119
  OuterSyntax.command "invoke"
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
   120
    "schematic invocation of locale expression in proof context"
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
   121
    (K.tag_proof K.prf_goal)
22101
6d13239d5f52 moved parts of OuterParse to SpecParse;
wenzelm
parents: 21584
diff changeset
   122
    (SpecParse.opt_thm_name ":" -- SpecParse.locale_expr -- SpecParse.locale_insts -- P.for_fixes
22658
263d42253f53 Experimental interpretation code for definitions.
ballarin
parents: 22568
diff changeset
   123
      >> (fn (((name, expr), (insts, _)), fixes) =>
19849
d96a15bb2d88 actually invoke result elements;
wenzelm
parents: 19813
diff changeset
   124
          Toplevel.print o Toplevel.proof' (invoke name expr insts fixes)));
19813
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
   125
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
   126
val _ = OuterSyntax.add_parsers [invokeP];
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
   127
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
   128
end;
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
   129
b051be997d50 Schematic invocation of locale expression in proof context.
wenzelm
parents:
diff changeset
   130
end;