src/Pure/more_thm.ML
author wenzelm
Wed, 31 Dec 2008 15:30:10 +0100
changeset 29269 5c25a2012975
parent 28965 1de908189869
child 29579 cb520b766e00
permissions -rw-r--r--
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML); tuned signature of structure Term;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
22362
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/more_thm.ML
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
     3
22907
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
     4
Further operations on type ctyp/cterm/thm, outside the inference kernel.
22362
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
     5
*)
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
     6
23169
37091da05d8e moved aconvc to more_thm.ML;
wenzelm
parents: 22907
diff changeset
     7
infix aconvc;
37091da05d8e moved aconvc to more_thm.ML;
wenzelm
parents: 22907
diff changeset
     8
22362
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
     9
signature THM =
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
    10
sig
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
    11
  include THM
24948
wenzelm
parents: 24775
diff changeset
    12
  val aconvc: cterm * cterm -> bool
23491
c13ca04303de added reasonably efficient add_cterm_frees;
wenzelm
parents: 23170
diff changeset
    13
  val add_cterm_frees: cterm -> cterm list -> cterm list
22907
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
    14
  val mk_binop: cterm -> cterm -> cterm -> cterm
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
    15
  val dest_binop: cterm -> cterm * cterm
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
    16
  val dest_implies: cterm -> cterm * cterm
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
    17
  val dest_equals: cterm -> cterm * cterm
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
    18
  val dest_equals_lhs: cterm -> cterm
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
    19
  val dest_equals_rhs: cterm -> cterm
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
    20
  val lhs_of: thm -> cterm
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
    21
  val rhs_of: thm -> cterm
22362
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
    22
  val thm_ord: thm * thm -> order
23599
d889725b0d8a added is_reflexive;
wenzelm
parents: 23491
diff changeset
    23
  val is_reflexive: thm -> bool
22362
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
    24
  val eq_thm: thm * thm -> bool
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
    25
  val eq_thms: thm list * thm list -> bool
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
    26
  val eq_thm_thy: thm * thm -> bool
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
    27
  val eq_thm_prop: thm * thm -> bool
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
    28
  val equiv_thm: thm * thm -> bool
28621
a60164e8fff0 added check_shyps, which reject pending sort hypotheses;
wenzelm
parents: 28116
diff changeset
    29
  val check_shyps: sort list -> thm -> thm
24048
a12b4faff474 moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents: 23599
diff changeset
    30
  val is_dummy: thm -> bool
22695
17073e9b94f2 moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents: 22682
diff changeset
    31
  val plain_prop_of: thm -> term
17073e9b94f2 moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents: 22682
diff changeset
    32
  val fold_terms: (term -> 'a -> 'a) -> thm -> 'a -> 'a
24048
a12b4faff474 moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents: 23599
diff changeset
    33
  val add_thm: thm -> thm list -> thm list
a12b4faff474 moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents: 23599
diff changeset
    34
  val del_thm: thm -> thm list -> thm list
a12b4faff474 moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents: 23599
diff changeset
    35
  val merge_thms: thm list * thm list -> thm list
27866
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
    36
  val elim_implies: thm -> thm -> thm
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
    37
  val forall_elim_var: int -> thm -> thm
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
    38
  val forall_elim_vars: int -> thm -> thm
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
    39
  val unvarify: thm -> thm
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
    40
  val close_derivation: thm -> thm
28116
cd2547ab0696 simplified add_axiom: no hyps;
wenzelm
parents: 28017
diff changeset
    41
  val add_axiom: bstring * term -> theory -> thm * theory
27866
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
    42
  val add_def: bool -> bool -> bstring * term -> theory -> thm * theory
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
    43
  val rule_attribute: (Context.generic -> thm -> thm) -> attribute
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
    44
  val declaration_attribute: (thm -> Context.generic -> Context.generic) -> attribute
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
    45
  val theory_attributes: attribute list -> theory * thm -> theory * thm
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
    46
  val proof_attributes: attribute list -> Proof.context * thm -> Proof.context * thm
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
    47
  val no_attributes: 'a -> 'a * 'b list
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
    48
  val simple_fact: 'a -> ('a * 'b list) list
28017
4919bd124a58 type Properties.T;
wenzelm
parents: 27866
diff changeset
    49
  val tag_rule: Properties.property -> thm -> thm
27866
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
    50
  val untag_rule: string -> thm -> thm
28017
4919bd124a58 type Properties.T;
wenzelm
parents: 27866
diff changeset
    51
  val tag: Properties.property -> attribute
27866
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
    52
  val untag: string -> attribute
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
    53
  val position_of: thm -> Position.T
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
    54
  val default_position: Position.T -> thm -> thm
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
    55
  val default_position_of: thm -> thm -> thm
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
    56
  val has_name_hint: thm -> bool
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
    57
  val get_name_hint: thm -> string
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
    58
  val put_name_hint: string -> thm -> thm
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
    59
  val get_group: thm -> string option
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
    60
  val put_group: string -> thm -> thm
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
    61
  val group: string -> attribute
22362
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
    62
  val axiomK: string
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
    63
  val assumptionK: string
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
    64
  val definitionK: string
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
    65
  val theoremK: string
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
    66
  val lemmaK: string
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
    67
  val corollaryK: string
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
    68
  val internalK: string
27866
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
    69
  val has_kind: thm -> bool
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
    70
  val get_kind: thm -> string
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
    71
  val kind_rule: string -> thm -> thm
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
    72
  val kind: string -> attribute
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
    73
  val kind_internal: attribute
28017
4919bd124a58 type Properties.T;
wenzelm
parents: 27866
diff changeset
    74
  val has_internal: Properties.property list -> bool
27866
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
    75
  val is_internal: thm -> bool
22362
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
    76
end;
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
    77
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
    78
structure Thm: THM =
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
    79
struct
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
    80
22695
17073e9b94f2 moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents: 22682
diff changeset
    81
(** basic operations **)
22362
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
    82
23491
c13ca04303de added reasonably efficient add_cterm_frees;
wenzelm
parents: 23170
diff changeset
    83
(* collecting cterms *)
c13ca04303de added reasonably efficient add_cterm_frees;
wenzelm
parents: 23170
diff changeset
    84
c13ca04303de added reasonably efficient add_cterm_frees;
wenzelm
parents: 23170
diff changeset
    85
val op aconvc = op aconv o pairself Thm.term_of;
c13ca04303de added reasonably efficient add_cterm_frees;
wenzelm
parents: 23170
diff changeset
    86
c13ca04303de added reasonably efficient add_cterm_frees;
wenzelm
parents: 23170
diff changeset
    87
fun add_cterm_frees ct =
c13ca04303de added reasonably efficient add_cterm_frees;
wenzelm
parents: 23170
diff changeset
    88
  let
c13ca04303de added reasonably efficient add_cterm_frees;
wenzelm
parents: 23170
diff changeset
    89
    val cert = Thm.cterm_of (Thm.theory_of_cterm ct);
c13ca04303de added reasonably efficient add_cterm_frees;
wenzelm
parents: 23170
diff changeset
    90
    val t = Thm.term_of ct;
c13ca04303de added reasonably efficient add_cterm_frees;
wenzelm
parents: 23170
diff changeset
    91
  in Term.fold_aterms (fn v as Free _ => insert (op aconvc) (cert v) | _ => I) t end;
c13ca04303de added reasonably efficient add_cterm_frees;
wenzelm
parents: 23170
diff changeset
    92
c13ca04303de added reasonably efficient add_cterm_frees;
wenzelm
parents: 23170
diff changeset
    93
22907
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
    94
(* cterm constructors and destructors *)
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
    95
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
    96
fun mk_binop c a b = Thm.capply (Thm.capply c a) b;
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
    97
fun dest_binop ct = (Thm.dest_arg1 ct, Thm.dest_arg ct);
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
    98
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
    99
fun dest_implies ct =
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
   100
  (case Thm.term_of ct of
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
   101
    Const ("==>", _) $ _ $ _ => dest_binop ct
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
   102
  | _ => raise TERM ("dest_implies", [Thm.term_of ct]));
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
   103
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
   104
fun dest_equals ct =
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
   105
  (case Thm.term_of ct of
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
   106
    Const ("==", _) $ _ $ _ => dest_binop ct
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
   107
  | _ => raise TERM ("dest_equals", [Thm.term_of ct]));
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
   108
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
   109
fun dest_equals_lhs ct =
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
   110
  (case Thm.term_of ct of
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
   111
    Const ("==", _) $ _ $ _ => Thm.dest_arg1 ct
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
   112
  | _ => raise TERM ("dest_equals_lhs", [Thm.term_of ct]));
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
   113
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
   114
fun dest_equals_rhs ct =
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
   115
  (case Thm.term_of ct of
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
   116
    Const ("==", _) $ _ $ _ => Thm.dest_arg ct
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
   117
  | _ => raise TERM ("dest_equals_rhs", [Thm.term_of ct]));
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
   118
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
   119
val lhs_of = dest_equals_lhs o Thm.cprop_of;
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
   120
val rhs_of = dest_equals_rhs o Thm.cprop_of;
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
   121
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
   122
dccd0763ae37 added destructors from drule.ML;
wenzelm
parents: 22695
diff changeset
   123
(* thm order: ignores theory context! *)
22682
92448396c9d9 added read_def_cterms, read_cterm (from thm.ML);
wenzelm
parents: 22378
diff changeset
   124
22362
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
   125
fun thm_ord (th1, th2) =
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
   126
  let
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
   127
    val {shyps = shyps1, hyps = hyps1, tpairs = tpairs1, prop = prop1, ...} = Thm.rep_thm th1;
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
   128
    val {shyps = shyps2, hyps = hyps2, tpairs = tpairs2, prop = prop2, ...} = Thm.rep_thm th2;
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
   129
  in
29269
5c25a2012975 moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
wenzelm
parents: 28965
diff changeset
   130
    (case TermOrd.fast_term_ord (prop1, prop2) of
22362
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
   131
      EQUAL =>
29269
5c25a2012975 moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
wenzelm
parents: 28965
diff changeset
   132
        (case list_ord (prod_ord TermOrd.fast_term_ord TermOrd.fast_term_ord) (tpairs1, tpairs2) of
22362
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
   133
          EQUAL =>
29269
5c25a2012975 moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
wenzelm
parents: 28965
diff changeset
   134
            (case list_ord TermOrd.fast_term_ord (hyps1, hyps2) of
5c25a2012975 moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
wenzelm
parents: 28965
diff changeset
   135
              EQUAL => list_ord TermOrd.sort_ord (shyps1, shyps2)
22362
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
   136
            | ord => ord)
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
   137
        | ord => ord)
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
   138
    | ord => ord)
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
   139
  end;
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
   140
22682
92448396c9d9 added read_def_cterms, read_cterm (from thm.ML);
wenzelm
parents: 22378
diff changeset
   141
92448396c9d9 added read_def_cterms, read_cterm (from thm.ML);
wenzelm
parents: 22378
diff changeset
   142
(* equality *)
92448396c9d9 added read_def_cterms, read_cterm (from thm.ML);
wenzelm
parents: 22378
diff changeset
   143
23599
d889725b0d8a added is_reflexive;
wenzelm
parents: 23491
diff changeset
   144
fun is_reflexive th = op aconv (Logic.dest_equals (Thm.prop_of th))
d889725b0d8a added is_reflexive;
wenzelm
parents: 23491
diff changeset
   145
  handle TERM _ => false;
d889725b0d8a added is_reflexive;
wenzelm
parents: 23491
diff changeset
   146
22362
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
   147
fun eq_thm ths =
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
   148
  Context.joinable (pairself Thm.theory_of_thm ths) andalso
26653
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   149
  is_equal (thm_ord ths);
22362
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
   150
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
   151
val eq_thms = eq_list eq_thm;
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
   152
26665
2e363edf7578 Theory.eq_thy;
wenzelm
parents: 26653
diff changeset
   153
val eq_thm_thy = Theory.eq_thy o pairself Thm.theory_of_thm;
22362
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
   154
val eq_thm_prop = op aconv o pairself Thm.full_prop_of;
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
   155
22682
92448396c9d9 added read_def_cterms, read_cterm (from thm.ML);
wenzelm
parents: 22378
diff changeset
   156
92448396c9d9 added read_def_cterms, read_cterm (from thm.ML);
wenzelm
parents: 22378
diff changeset
   157
(* pattern equivalence *)
92448396c9d9 added read_def_cterms, read_cterm (from thm.ML);
wenzelm
parents: 22378
diff changeset
   158
22362
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
   159
fun equiv_thm ths =
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
   160
  Pattern.equiv (Theory.merge (pairself Thm.theory_of_thm ths)) (pairself Thm.full_prop_of ths);
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
   161
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
   162
28621
a60164e8fff0 added check_shyps, which reject pending sort hypotheses;
wenzelm
parents: 28116
diff changeset
   163
(* sort hypotheses *)
a60164e8fff0 added check_shyps, which reject pending sort hypotheses;
wenzelm
parents: 28116
diff changeset
   164
a60164e8fff0 added check_shyps, which reject pending sort hypotheses;
wenzelm
parents: 28116
diff changeset
   165
fun check_shyps sorts raw_th =
a60164e8fff0 added check_shyps, which reject pending sort hypotheses;
wenzelm
parents: 28116
diff changeset
   166
  let
a60164e8fff0 added check_shyps, which reject pending sort hypotheses;
wenzelm
parents: 28116
diff changeset
   167
    val th = Thm.strip_shyps raw_th;
a60164e8fff0 added check_shyps, which reject pending sort hypotheses;
wenzelm
parents: 28116
diff changeset
   168
    val prt_sort = Syntax.pretty_sort_global (Thm.theory_of_thm th);
a60164e8fff0 added check_shyps, which reject pending sort hypotheses;
wenzelm
parents: 28116
diff changeset
   169
    val pending = Sorts.subtract sorts (Thm.extra_shyps th);
a60164e8fff0 added check_shyps, which reject pending sort hypotheses;
wenzelm
parents: 28116
diff changeset
   170
  in
a60164e8fff0 added check_shyps, which reject pending sort hypotheses;
wenzelm
parents: 28116
diff changeset
   171
    if null pending then th
a60164e8fff0 added check_shyps, which reject pending sort hypotheses;
wenzelm
parents: 28116
diff changeset
   172
    else error (Pretty.string_of (Pretty.block (Pretty.str "Pending sort hypotheses:" ::
a60164e8fff0 added check_shyps, which reject pending sort hypotheses;
wenzelm
parents: 28116
diff changeset
   173
      Pretty.brk 1 :: Pretty.commas (map prt_sort pending))))
a60164e8fff0 added check_shyps, which reject pending sort hypotheses;
wenzelm
parents: 28116
diff changeset
   174
  end;
a60164e8fff0 added check_shyps, which reject pending sort hypotheses;
wenzelm
parents: 28116
diff changeset
   175
a60164e8fff0 added check_shyps, which reject pending sort hypotheses;
wenzelm
parents: 28116
diff changeset
   176
22695
17073e9b94f2 moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents: 22682
diff changeset
   177
(* misc operations *)
17073e9b94f2 moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents: 22682
diff changeset
   178
24048
a12b4faff474 moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents: 23599
diff changeset
   179
fun is_dummy thm =
a12b4faff474 moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents: 23599
diff changeset
   180
  (case try Logic.dest_term (Thm.concl_of thm) of
a12b4faff474 moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents: 23599
diff changeset
   181
    NONE => false
a12b4faff474 moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents: 23599
diff changeset
   182
  | SOME t => Term.is_dummy_pattern t);
a12b4faff474 moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents: 23599
diff changeset
   183
22695
17073e9b94f2 moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents: 22682
diff changeset
   184
fun plain_prop_of raw_thm =
17073e9b94f2 moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents: 22682
diff changeset
   185
  let
17073e9b94f2 moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents: 22682
diff changeset
   186
    val thm = Thm.strip_shyps raw_thm;
17073e9b94f2 moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents: 22682
diff changeset
   187
    fun err msg = raise THM ("plain_prop_of: " ^ msg, 0, [thm]);
17073e9b94f2 moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents: 22682
diff changeset
   188
    val {hyps, prop, tpairs, ...} = Thm.rep_thm thm;
17073e9b94f2 moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents: 22682
diff changeset
   189
  in
17073e9b94f2 moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents: 22682
diff changeset
   190
    if not (null hyps) then
17073e9b94f2 moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents: 22682
diff changeset
   191
      err "theorem may not contain hypotheses"
17073e9b94f2 moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents: 22682
diff changeset
   192
    else if not (null (Thm.extra_shyps thm)) then
17073e9b94f2 moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents: 22682
diff changeset
   193
      err "theorem may not contain sort hypotheses"
17073e9b94f2 moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents: 22682
diff changeset
   194
    else if not (null tpairs) then
17073e9b94f2 moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents: 22682
diff changeset
   195
      err "theorem may not contain flex-flex pairs"
17073e9b94f2 moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents: 22682
diff changeset
   196
    else prop
17073e9b94f2 moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents: 22682
diff changeset
   197
  end;
17073e9b94f2 moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents: 22682
diff changeset
   198
17073e9b94f2 moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents: 22682
diff changeset
   199
fun fold_terms f th =
17073e9b94f2 moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents: 22682
diff changeset
   200
  let val {tpairs, prop, hyps, ...} = Thm.rep_thm th
17073e9b94f2 moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents: 22682
diff changeset
   201
  in fold (fn (t, u) => f t #> f u) tpairs #> f prop #> fold f hyps end;
17073e9b94f2 moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents: 22682
diff changeset
   202
17073e9b94f2 moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents: 22682
diff changeset
   203
24048
a12b4faff474 moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents: 23599
diff changeset
   204
(* lists of theorems in canonical order *)
a12b4faff474 moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents: 23599
diff changeset
   205
a12b4faff474 moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents: 23599
diff changeset
   206
val add_thm = update eq_thm_prop;
a12b4faff474 moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents: 23599
diff changeset
   207
val del_thm = remove eq_thm_prop;
a12b4faff474 moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents: 23599
diff changeset
   208
val merge_thms = merge eq_thm_prop;
a12b4faff474 moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents: 23599
diff changeset
   209
a12b4faff474 moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents: 23599
diff changeset
   210
22682
92448396c9d9 added read_def_cterms, read_cterm (from thm.ML);
wenzelm
parents: 22378
diff changeset
   211
24980
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   212
(** basic derived rules **)
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   213
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   214
(*Elimination of implication
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   215
  A    A ==> B
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   216
  ------------
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   217
        B
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   218
*)
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   219
fun elim_implies thA thAB = Thm.implies_elim thAB thA;
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   220
26653
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   221
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   222
(* forall_elim_var(s) *)
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   223
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   224
local
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   225
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   226
fun forall_elim_vars_aux strip_vars i th =
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   227
  let
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   228
    val thy = Thm.theory_of_thm th;
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   229
    val {tpairs, prop, ...} = Thm.rep_thm th;
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   230
    val add_used = Term.fold_aterms
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   231
      (fn Var ((x, j), _) => if i = j then insert (op =) x else I | _ => I);
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   232
    val used = fold (fn (t, u) => add_used t o add_used u) tpairs (add_used prop []);
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   233
    val vars = strip_vars prop;
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   234
    val cvars = (Name.variant_list used (map #1 vars), vars)
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   235
      |> ListPair.map (fn (x, (_, T)) => Thm.cterm_of thy (Var ((x, i), T)));
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   236
  in fold Thm.forall_elim cvars th end;
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   237
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   238
in
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   239
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   240
val forall_elim_vars = forall_elim_vars_aux Term.strip_all_vars;
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   241
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   242
fun forall_elim_var i th = forall_elim_vars_aux
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   243
  (fn Const ("all", _) $ Abs (a, T, _) => [(a, T)]
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   244
  | _ => raise THM ("forall_elim_vars", i, [th])) i th;
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   245
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   246
end;
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   247
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   248
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   249
(* unvarify: global schematic variables *)
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   250
24980
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   251
fun unvarify th =
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   252
  let
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   253
    val thy = Thm.theory_of_thm th;
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   254
    val cert = Thm.cterm_of thy;
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   255
    val certT = Thm.ctyp_of thy;
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   256
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   257
    val prop = Thm.full_prop_of th;
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   258
    val _ = map Logic.unvarify (prop :: Thm.hyps_of th)
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   259
      handle TERM (msg, _) => raise THM (msg, 0, [th]);
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   260
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   261
    val instT0 = rev (Term.add_tvars prop []) |> map (fn v as ((a, _), S) => (v, TFree (a, S)));
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   262
    val instT = map (fn (v, T) => (certT (TVar v), certT T)) instT0;
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   263
    val inst = rev (Term.add_vars prop []) |> map (fn ((a, i), T) =>
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   264
      let val T' = TermSubst.instantiateT instT0 T
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   265
      in (cert (Var ((a, i), T')), cert (Free ((a, T')))) end);
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   266
  in Thm.instantiate (instT, inst) th end;
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   267
26653
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   268
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   269
(* close_derivation *)
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26628
diff changeset
   270
26628
63306cb94313 replaced Drule.close_derivation/Goal.close_result by Thm.close_derivation (removed obsolete compression);
wenzelm
parents: 25518
diff changeset
   271
fun close_derivation thm =
63306cb94313 replaced Drule.close_derivation/Goal.close_result by Thm.close_derivation (removed obsolete compression);
wenzelm
parents: 25518
diff changeset
   272
  if Thm.get_name thm = "" then Thm.put_name "" thm
63306cb94313 replaced Drule.close_derivation/Goal.close_result by Thm.close_derivation (removed obsolete compression);
wenzelm
parents: 25518
diff changeset
   273
  else thm;
63306cb94313 replaced Drule.close_derivation/Goal.close_result by Thm.close_derivation (removed obsolete compression);
wenzelm
parents: 25518
diff changeset
   274
24980
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   275
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   276
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   277
(** specification primitives **)
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   278
28116
cd2547ab0696 simplified add_axiom: no hyps;
wenzelm
parents: 28017
diff changeset
   279
fun add_axiom (name, prop) thy =
24980
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   280
  let
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   281
    val name' = if name = "" then "axiom_" ^ serial_string () else name;
28116
cd2547ab0696 simplified add_axiom: no hyps;
wenzelm
parents: 28017
diff changeset
   282
    val thy' = thy |> Theory.add_axioms_i [(name', prop)];
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28674
diff changeset
   283
    val axm = unvarify (Thm.axiom thy' (Sign.full_bname thy' name'));
28116
cd2547ab0696 simplified add_axiom: no hyps;
wenzelm
parents: 28017
diff changeset
   284
  in (axm, thy') end;
24980
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   285
25518
00d5cc16e891 interface for unchecked definitions
haftmann
parents: 24980
diff changeset
   286
fun add_def unchecked overloaded (name, prop) thy =
24980
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   287
  let
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   288
    val tfrees = rev (map TFree (Term.add_tfrees prop []));
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   289
    val tfrees' = map (fn a => TFree (a, [])) (Name.invents Name.context Name.aT (length tfrees));
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   290
    val strip_sorts = tfrees ~~ tfrees';
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   291
    val recover_sorts = map (pairself (Thm.ctyp_of thy o Logic.varifyT)) (tfrees' ~~ tfrees);
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   292
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   293
    val prop' = Term.map_types (Term.map_atyps (perhaps (AList.lookup (op =) strip_sorts))) prop;
25518
00d5cc16e891 interface for unchecked definitions
haftmann
parents: 24980
diff changeset
   294
    val thy' = Theory.add_defs_i unchecked overloaded [(name, prop')] thy;
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28674
diff changeset
   295
    val axm' = Thm.axiom thy' (Sign.full_bname thy' name);
24980
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   296
    val thm = unvarify (Thm.instantiate (recover_sorts, []) axm');
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   297
  in (thm, thy') end;
16a74cfca971 added elim_implies (more convenient argument order);
wenzelm
parents: 24948
diff changeset
   298
27866
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   299
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   300
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   301
(** attributes **)
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   302
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   303
fun rule_attribute f (x, th) = (x, f x th);
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   304
fun declaration_attribute f (x, th) = (f th x, th);
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   305
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   306
fun apply_attributes mk dest =
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   307
  let
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   308
    fun app [] = I
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   309
      | app ((f: attribute) :: fs) = fn (x, th) => f (mk x, th) |>> dest |> app fs;
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   310
  in app end;
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   311
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   312
val theory_attributes = apply_attributes Context.Theory Context.the_theory;
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   313
val proof_attributes = apply_attributes Context.Proof Context.the_proof;
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   314
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   315
fun no_attributes x = (x, []);
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   316
fun simple_fact x = [(x, [])];
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   317
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   318
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   319
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   320
(*** theorem tags ***)
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   321
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   322
(* add / delete tags *)
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   323
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   324
fun tag_rule tg = Thm.map_tags (insert (op =) tg);
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   325
fun untag_rule s = Thm.map_tags (filter_out (fn (s', _) => s = s'));
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   326
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   327
fun tag tg x = rule_attribute (K (tag_rule tg)) x;
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   328
fun untag s x = rule_attribute (K (untag_rule s)) x;
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   329
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   330
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   331
(* position *)
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   332
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   333
val position_of = Position.of_properties o Thm.get_tags;
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   334
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   335
fun default_position pos = Thm.map_tags (Position.default_properties pos);
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   336
val default_position_of = default_position o position_of;
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   337
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   338
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   339
(* unofficial theorem names *)
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   340
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   341
fun the_name_hint thm = the (AList.lookup (op =) (Thm.get_tags thm) Markup.nameN);
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   342
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   343
val has_name_hint = can the_name_hint;
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   344
val get_name_hint = the_default "??.unknown" o try the_name_hint;
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   345
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   346
fun put_name_hint name = untag_rule Markup.nameN #> tag_rule (Markup.nameN, name);
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   347
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   348
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   349
(* theorem groups *)
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   350
28017
4919bd124a58 type Properties.T;
wenzelm
parents: 27866
diff changeset
   351
fun get_group thm = Properties.get (Thm.get_tags thm) Markup.groupN;
27866
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   352
28017
4919bd124a58 type Properties.T;
wenzelm
parents: 27866
diff changeset
   353
fun put_group name = if name = "" then I else Thm.map_tags (Properties.put (Markup.groupN, name));
27866
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   354
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   355
fun group name = rule_attribute (K (put_group name));
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   356
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   357
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   358
(* theorem kinds *)
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   359
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   360
val axiomK = "axiom";
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   361
val assumptionK = "assumption";
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   362
val definitionK = "definition";
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   363
val theoremK = "theorem";
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   364
val lemmaK = "lemma";
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   365
val corollaryK = "corollary";
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   366
val internalK = Markup.internalK;
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   367
28017
4919bd124a58 type Properties.T;
wenzelm
parents: 27866
diff changeset
   368
fun the_kind thm = the (Properties.get (Thm.get_tags thm) Markup.kindN);
27866
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   369
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   370
val has_kind = can the_kind;
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   371
val get_kind = the_default "" o try the_kind;
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   372
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   373
fun kind_rule k = tag_rule (Markup.kindN, k) o untag_rule Markup.kindN;
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   374
fun kind k x = if k = "" then x else rule_attribute (K (kind_rule k)) x;
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   375
fun kind_internal x = kind internalK x;
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   376
fun has_internal tags = exists (fn tg => tg = (Markup.kindN, internalK)) tags;
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   377
val is_internal = has_internal o Thm.get_tags;
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   378
c721ea6e0eb4 moved basic thm operations from structure PureThy to Thm;
wenzelm
parents: 27255
diff changeset
   379
22362
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
   380
open Thm;
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
   381
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
   382
end;
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
   383
23170
94e9413bd7fc made aconvc pervasive;
wenzelm
parents: 23169
diff changeset
   384
val op aconvc = Thm.aconvc;
94e9413bd7fc made aconvc pervasive;
wenzelm
parents: 23169
diff changeset
   385
22362
6470ce514b6e Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff changeset
   386
structure Thmtab = TableFun(type key = thm val ord = Thm.thm_ord);