src/Pure/Isar/theory_target.ML
author haftmann
Mon, 05 Jan 2009 15:36:24 +0100
changeset 29358 efdfe5dfe008
parent 29252 ea97aa6aeba2
child 29360 a5be60c3674e
permissions -rw-r--r--
rearranged target theories
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
20894
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Isar/theory_target.ML
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
     3
25542
haftmann
parents: 25519
diff changeset
     4
Common theory/locale/class/instantiation/overloading targets.
20894
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
     5
*)
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
     6
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
     7
signature THEORY_TARGET =
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
     8
sig
29228
40b3630b0deb Theory target distinguishes old and new locales.
ballarin
parents: 29223
diff changeset
     9
  val peek: local_theory -> {target: string, new_locale: bool, is_locale: bool,
25864
11f531354852 explicit type variables for instantiation
haftmann
parents: 25684
diff changeset
    10
    is_class: bool, instantiation: string list * (string * sort) list * sort,
11f531354852 explicit type variables for instantiation
haftmann
parents: 25684
diff changeset
    11
    overloading: (string * (string * typ) * bool) list}
25269
f9090ae5cec9 clarified theory target interface
haftmann
parents: 25240
diff changeset
    12
  val init: string option -> theory -> local_theory
24935
a033971c63a0 removed LocalTheory.defs/target_morphism operations;
wenzelm
parents: 24914
diff changeset
    13
  val begin: string -> Proof.context -> local_theory
25291
870455627720 misc cleanup of init functions;
wenzelm
parents: 25269
diff changeset
    14
  val context: xstring -> theory -> local_theory
25864
11f531354852 explicit type variables for instantiation
haftmann
parents: 25684
diff changeset
    15
  val instantiation: string list * (string * sort) list * sort -> theory -> local_theory
29358
efdfe5dfe008 rearranged target theories
haftmann
parents: 29252
diff changeset
    16
  val instantiation_cmd: xstring list * sort * xstring -> theory -> local_theory
25864
11f531354852 explicit type variables for instantiation
haftmann
parents: 25684
diff changeset
    17
  val overloading: (string * (string * typ) * bool) list -> theory -> local_theory
26049
8186c03194ed overloading: reduced code redundancy, no xstrings here;
wenzelm
parents: 25984
diff changeset
    18
  val overloading_cmd: (string * string * bool) list -> theory -> local_theory
20894
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    19
end;
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    20
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    21
structure TheoryTarget: THEORY_TARGET =
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    22
struct
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    23
28834
66d31ca2c5af Code for switching to new locales.
ballarin
parents: 28822
diff changeset
    24
(* new locales *)
66d31ca2c5af Code for switching to new locales.
ballarin
parents: 28822
diff changeset
    25
29228
40b3630b0deb Theory target distinguishes old and new locales.
ballarin
parents: 29223
diff changeset
    26
fun locale_extern new_locale x = 
29249
4dc278c8dc59 All logics ported to new locales.
ballarin
parents: 29228
diff changeset
    27
  if new_locale then NewLocale.extern x else Locale.extern x;
29228
40b3630b0deb Theory target distinguishes old and new locales.
ballarin
parents: 29223
diff changeset
    28
fun locale_add_type_syntax new_locale x =
29249
4dc278c8dc59 All logics ported to new locales.
ballarin
parents: 29228
diff changeset
    29
  if new_locale then NewLocale.add_type_syntax x else Locale.add_type_syntax x;
29228
40b3630b0deb Theory target distinguishes old and new locales.
ballarin
parents: 29223
diff changeset
    30
fun locale_add_term_syntax new_locale x =
29249
4dc278c8dc59 All logics ported to new locales.
ballarin
parents: 29228
diff changeset
    31
  if new_locale then NewLocale.add_term_syntax x else Locale.add_term_syntax x;
29228
40b3630b0deb Theory target distinguishes old and new locales.
ballarin
parents: 29223
diff changeset
    32
fun locale_add_declaration new_locale x =
29249
4dc278c8dc59 All logics ported to new locales.
ballarin
parents: 29228
diff changeset
    33
  if new_locale then NewLocale.add_declaration x else Locale.add_declaration x;
29228
40b3630b0deb Theory target distinguishes old and new locales.
ballarin
parents: 29223
diff changeset
    34
fun locale_add_thmss new_locale x =
29249
4dc278c8dc59 All logics ported to new locales.
ballarin
parents: 29228
diff changeset
    35
  if new_locale then NewLocale.add_thmss x else Locale.add_thmss x;
29228
40b3630b0deb Theory target distinguishes old and new locales.
ballarin
parents: 29223
diff changeset
    36
fun locale_init new_locale x =
29249
4dc278c8dc59 All logics ported to new locales.
ballarin
parents: 29228
diff changeset
    37
  if new_locale then NewLocale.init x else Locale.init x;
29228
40b3630b0deb Theory target distinguishes old and new locales.
ballarin
parents: 29223
diff changeset
    38
fun locale_intern new_locale x =
29249
4dc278c8dc59 All logics ported to new locales.
ballarin
parents: 29228
diff changeset
    39
  if new_locale then NewLocale.intern x else Locale.intern x;
28834
66d31ca2c5af Code for switching to new locales.
ballarin
parents: 28822
diff changeset
    40
21006
ac2732072403 added peek;
wenzelm
parents: 20984
diff changeset
    41
(* context data *)
ac2732072403 added peek;
wenzelm
parents: 20984
diff changeset
    42
29228
40b3630b0deb Theory target distinguishes old and new locales.
ballarin
parents: 29223
diff changeset
    43
datatype target = Target of {target: string, new_locale: bool, is_locale: bool,
25864
11f531354852 explicit type variables for instantiation
haftmann
parents: 25684
diff changeset
    44
  is_class: bool, instantiation: string list * (string * sort) list * sort,
11f531354852 explicit type variables for instantiation
haftmann
parents: 25684
diff changeset
    45
  overloading: (string * (string * typ) * bool) list};
25012
448af76a1ba3 pass explicit target record -- more informative peek operation;
wenzelm
parents: 25002
diff changeset
    46
29228
40b3630b0deb Theory target distinguishes old and new locales.
ballarin
parents: 29223
diff changeset
    47
fun make_target target new_locale is_locale is_class instantiation overloading =
40b3630b0deb Theory target distinguishes old and new locales.
ballarin
parents: 29223
diff changeset
    48
  Target {target = target, new_locale = new_locale, is_locale = is_locale,
25519
8570745cb40b overloading target
haftmann
parents: 25514
diff changeset
    49
    is_class = is_class, instantiation = instantiation, overloading = overloading};
25291
870455627720 misc cleanup of init functions;
wenzelm
parents: 25269
diff changeset
    50
29228
40b3630b0deb Theory target distinguishes old and new locales.
ballarin
parents: 29223
diff changeset
    51
val global_target = make_target "" false false false ([], [], []) [];
25012
448af76a1ba3 pass explicit target record -- more informative peek operation;
wenzelm
parents: 25002
diff changeset
    52
21006
ac2732072403 added peek;
wenzelm
parents: 20984
diff changeset
    53
structure Data = ProofDataFun
ac2732072403 added peek;
wenzelm
parents: 20984
diff changeset
    54
(
25012
448af76a1ba3 pass explicit target record -- more informative peek operation;
wenzelm
parents: 25002
diff changeset
    55
  type T = target;
25291
870455627720 misc cleanup of init functions;
wenzelm
parents: 25269
diff changeset
    56
  fun init _ = global_target;
21006
ac2732072403 added peek;
wenzelm
parents: 20984
diff changeset
    57
);
ac2732072403 added peek;
wenzelm
parents: 20984
diff changeset
    58
25012
448af76a1ba3 pass explicit target record -- more informative peek operation;
wenzelm
parents: 25002
diff changeset
    59
val peek = (fn Target args => args) o Data.get;
21006
ac2732072403 added peek;
wenzelm
parents: 20984
diff changeset
    60
ac2732072403 added peek;
wenzelm
parents: 20984
diff changeset
    61
20894
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    62
(* pretty *)
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    63
25607
779c79c36c5e pretty for instantiation and overloading
haftmann
parents: 25597
diff changeset
    64
fun pretty_thy ctxt target is_locale is_class =
20894
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    65
  let
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    66
    val thy = ProofContext.theory_of ctxt;
29223
e09c53289830 Conversion of HOL-Main and ZF to new locales.
ballarin
parents: 29006
diff changeset
    67
    val target_name = (if is_class then "class " else "locale ") ^ locale_extern is_class thy target;
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28941
diff changeset
    68
    val fixes = map (fn (x, T) => (Binding.name x, SOME T, NoSyn))
28094
wenzelm
parents: 28084
diff changeset
    69
      (#1 (ProofContext.inferred_fixes ctxt));
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28941
diff changeset
    70
    val assumes = map (fn A => (Attrib.empty_binding, [(Thm.term_of A, [])]))
28094
wenzelm
parents: 28084
diff changeset
    71
      (Assumption.assms_of ctxt);
20894
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    72
    val elems =
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    73
      (if null fixes then [] else [Element.Fixes fixes]) @
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    74
      (if null assumes then [] else [Element.Assumes assumes]);
26049
8186c03194ed overloading: reduced code redundancy, no xstrings here;
wenzelm
parents: 25984
diff changeset
    75
  in
8186c03194ed overloading: reduced code redundancy, no xstrings here;
wenzelm
parents: 25984
diff changeset
    76
    if target = "" then []
25607
779c79c36c5e pretty for instantiation and overloading
haftmann
parents: 25597
diff changeset
    77
    else if null elems then [Pretty.str target_name]
779c79c36c5e pretty for instantiation and overloading
haftmann
parents: 25597
diff changeset
    78
    else [Pretty.big_list (target_name ^ " =")
779c79c36c5e pretty for instantiation and overloading
haftmann
parents: 25597
diff changeset
    79
      (map (Pretty.chunks o Element.pretty_ctxt ctxt) elems)]
20894
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    80
  end;
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    81
29228
40b3630b0deb Theory target distinguishes old and new locales.
ballarin
parents: 29223
diff changeset
    82
fun pretty (Target {target, is_locale, is_class, instantiation, overloading, ...}) ctxt =
25607
779c79c36c5e pretty for instantiation and overloading
haftmann
parents: 25597
diff changeset
    83
  Pretty.block [Pretty.str "theory", Pretty.brk 1,
779c79c36c5e pretty for instantiation and overloading
haftmann
parents: 25597
diff changeset
    84
      Pretty.str (Context.theory_name (ProofContext.theory_of ctxt))] ::
779c79c36c5e pretty for instantiation and overloading
haftmann
parents: 25597
diff changeset
    85
    (if not (null overloading) then [Overloading.pretty ctxt]
29358
efdfe5dfe008 rearranged target theories
haftmann
parents: 29252
diff changeset
    86
     else if not (null (#1 instantiation)) then [Class_Target.pretty_instantiation ctxt]
26049
8186c03194ed overloading: reduced code redundancy, no xstrings here;
wenzelm
parents: 25984
diff changeset
    87
     else pretty_thy ctxt target is_locale is_class);
25607
779c79c36c5e pretty for instantiation and overloading
haftmann
parents: 25597
diff changeset
    88
20894
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    89
24838
1d1bddf87353 put declarations first
haftmann
parents: 24818
diff changeset
    90
(* target declarations *)
1d1bddf87353 put declarations first
haftmann
parents: 24818
diff changeset
    91
29228
40b3630b0deb Theory target distinguishes old and new locales.
ballarin
parents: 29223
diff changeset
    92
fun target_decl add (Target {target, new_locale, ...}) d lthy =
24838
1d1bddf87353 put declarations first
haftmann
parents: 24818
diff changeset
    93
  let
24935
a033971c63a0 removed LocalTheory.defs/target_morphism operations;
wenzelm
parents: 24914
diff changeset
    94
    val d' = Morphism.transform (LocalTheory.target_morphism lthy) d;
24838
1d1bddf87353 put declarations first
haftmann
parents: 24818
diff changeset
    95
    val d0 = Morphism.form d';
1d1bddf87353 put declarations first
haftmann
parents: 24818
diff changeset
    96
  in
25012
448af76a1ba3 pass explicit target record -- more informative peek operation;
wenzelm
parents: 25002
diff changeset
    97
    if target = "" then
24838
1d1bddf87353 put declarations first
haftmann
parents: 24818
diff changeset
    98
      lthy
1d1bddf87353 put declarations first
haftmann
parents: 24818
diff changeset
    99
      |> LocalTheory.theory (Context.theory_map d0)
1d1bddf87353 put declarations first
haftmann
parents: 24818
diff changeset
   100
      |> LocalTheory.target (Context.proof_map d0)
1d1bddf87353 put declarations first
haftmann
parents: 24818
diff changeset
   101
    else
1d1bddf87353 put declarations first
haftmann
parents: 24818
diff changeset
   102
      lthy
29228
40b3630b0deb Theory target distinguishes old and new locales.
ballarin
parents: 29223
diff changeset
   103
      |> LocalTheory.target (add new_locale target d')
24838
1d1bddf87353 put declarations first
haftmann
parents: 24818
diff changeset
   104
  end;
1d1bddf87353 put declarations first
haftmann
parents: 24818
diff changeset
   105
28834
66d31ca2c5af Code for switching to new locales.
ballarin
parents: 28822
diff changeset
   106
val type_syntax = target_decl locale_add_type_syntax;
66d31ca2c5af Code for switching to new locales.
ballarin
parents: 28822
diff changeset
   107
val term_syntax = target_decl locale_add_term_syntax;
66d31ca2c5af Code for switching to new locales.
ballarin
parents: 28822
diff changeset
   108
val declaration = target_decl locale_add_declaration;
24838
1d1bddf87353 put declarations first
haftmann
parents: 24818
diff changeset
   109
25105
c9f67836c4d8 internal tuning: class_target, fork_mixfix, declare_const (singleton), abbrev;
wenzelm
parents: 25096
diff changeset
   110
fun class_target (Target {target, ...}) f =
c9f67836c4d8 internal tuning: class_target, fork_mixfix, declare_const (singleton), abbrev;
wenzelm
parents: 25096
diff changeset
   111
  LocalTheory.raw_theory f #>
29358
efdfe5dfe008 rearranged target theories
haftmann
parents: 29252
diff changeset
   112
  LocalTheory.target (Class_Target.refresh_syntax target);
24838
1d1bddf87353 put declarations first
haftmann
parents: 24818
diff changeset
   113
21611
fc95ff1fe738 notes: proper import/export of proofs (still inactive);
wenzelm
parents: 21594
diff changeset
   114
20894
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   115
(* notes *)
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   116
21611
fc95ff1fe738 notes: proper import/export of proofs (still inactive);
wenzelm
parents: 21594
diff changeset
   117
fun import_export_proof ctxt (name, raw_th) =
21594
2859c94d67d4 reworked notes: towards proper import/export of proof terms;
wenzelm
parents: 21585
diff changeset
   118
  let
21611
fc95ff1fe738 notes: proper import/export of proofs (still inactive);
wenzelm
parents: 21594
diff changeset
   119
    val thy = ProofContext.theory_of ctxt;
fc95ff1fe738 notes: proper import/export of proofs (still inactive);
wenzelm
parents: 21594
diff changeset
   120
    val thy_ctxt = ProofContext.init thy;
fc95ff1fe738 notes: proper import/export of proofs (still inactive);
wenzelm
parents: 21594
diff changeset
   121
    val certT = Thm.ctyp_of thy;
fc95ff1fe738 notes: proper import/export of proofs (still inactive);
wenzelm
parents: 21594
diff changeset
   122
    val cert = Thm.cterm_of thy;
21594
2859c94d67d4 reworked notes: towards proper import/export of proof terms;
wenzelm
parents: 21585
diff changeset
   123
21611
fc95ff1fe738 notes: proper import/export of proofs (still inactive);
wenzelm
parents: 21594
diff changeset
   124
    (*export assumes/defines*)
fc95ff1fe738 notes: proper import/export of proofs (still inactive);
wenzelm
parents: 21594
diff changeset
   125
    val th = Goal.norm_result raw_th;
fc95ff1fe738 notes: proper import/export of proofs (still inactive);
wenzelm
parents: 21594
diff changeset
   126
    val (defs, th') = LocalDefs.export ctxt thy_ctxt th;
21708
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21691
diff changeset
   127
    val concl_conv = MetaSimplifier.rewrite true defs (Thm.cprop_of th);
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21691
diff changeset
   128
    val assms = map (MetaSimplifier.rewrite_rule defs o Thm.assume) (Assumption.assms_of ctxt);
21611
fc95ff1fe738 notes: proper import/export of proofs (still inactive);
wenzelm
parents: 21594
diff changeset
   129
    val nprems = Thm.nprems_of th' - Thm.nprems_of th;
fc95ff1fe738 notes: proper import/export of proofs (still inactive);
wenzelm
parents: 21594
diff changeset
   130
fc95ff1fe738 notes: proper import/export of proofs (still inactive);
wenzelm
parents: 21594
diff changeset
   131
    (*export fixes*)
22692
1e057a3f087d proper ProofContext.infer_types;
wenzelm
parents: 22673
diff changeset
   132
    val tfrees = map TFree (Thm.fold_terms Term.add_tfrees th' []);
1e057a3f087d proper ProofContext.infer_types;
wenzelm
parents: 22673
diff changeset
   133
    val frees = map Free (Thm.fold_terms Term.add_frees th' []);
21611
fc95ff1fe738 notes: proper import/export of proofs (still inactive);
wenzelm
parents: 21594
diff changeset
   134
    val (th'' :: vs) = (th' :: map (Drule.mk_term o cert) (map Logic.mk_type tfrees @ frees))
fc95ff1fe738 notes: proper import/export of proofs (still inactive);
wenzelm
parents: 21594
diff changeset
   135
      |> Variable.export ctxt thy_ctxt
fc95ff1fe738 notes: proper import/export of proofs (still inactive);
wenzelm
parents: 21594
diff changeset
   136
      |> Drule.zero_var_indexes_list;
fc95ff1fe738 notes: proper import/export of proofs (still inactive);
wenzelm
parents: 21594
diff changeset
   137
fc95ff1fe738 notes: proper import/export of proofs (still inactive);
wenzelm
parents: 21594
diff changeset
   138
    (*thm definition*)
28083
103d9282a946 explicit type Name.binding for higher-specification elements;
wenzelm
parents: 27690
diff changeset
   139
    val result = PureThy.name_thm true true Position.none name th'';
21611
fc95ff1fe738 notes: proper import/export of proofs (still inactive);
wenzelm
parents: 21594
diff changeset
   140
fc95ff1fe738 notes: proper import/export of proofs (still inactive);
wenzelm
parents: 21594
diff changeset
   141
    (*import fixes*)
fc95ff1fe738 notes: proper import/export of proofs (still inactive);
wenzelm
parents: 21594
diff changeset
   142
    val (tvars, vars) =
fc95ff1fe738 notes: proper import/export of proofs (still inactive);
wenzelm
parents: 21594
diff changeset
   143
      chop (length tfrees) (map (Thm.term_of o Drule.dest_term) vs)
fc95ff1fe738 notes: proper import/export of proofs (still inactive);
wenzelm
parents: 21594
diff changeset
   144
      |>> map Logic.dest_type;
fc95ff1fe738 notes: proper import/export of proofs (still inactive);
wenzelm
parents: 21594
diff changeset
   145
fc95ff1fe738 notes: proper import/export of proofs (still inactive);
wenzelm
parents: 21594
diff changeset
   146
    val instT = map_filter (fn (TVar v, T) => SOME (v, T) | _ => NONE) (tvars ~~ tfrees);
fc95ff1fe738 notes: proper import/export of proofs (still inactive);
wenzelm
parents: 21594
diff changeset
   147
    val inst = filter (is_Var o fst) (vars ~~ frees);
fc95ff1fe738 notes: proper import/export of proofs (still inactive);
wenzelm
parents: 21594
diff changeset
   148
    val cinstT = map (pairself certT o apfst TVar) instT;
fc95ff1fe738 notes: proper import/export of proofs (still inactive);
wenzelm
parents: 21594
diff changeset
   149
    val cinst = map (pairself (cert o Term.map_types (TermSubst.instantiateT instT))) inst;
fc95ff1fe738 notes: proper import/export of proofs (still inactive);
wenzelm
parents: 21594
diff changeset
   150
    val result' = Thm.instantiate (cinstT, cinst) result;
fc95ff1fe738 notes: proper import/export of proofs (still inactive);
wenzelm
parents: 21594
diff changeset