author | wenzelm |
Mon, 18 Aug 2014 15:46:27 +0200 | |
changeset 58001 | 934d85f14d1d |
parent 56245 | 84fc7dfa3cd4 |
child 59058 | a78612c67ec0 |
permissions | -rw-r--r-- |
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 | 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 | 7 |
infix aconvc; |
8 |
||
32842 | 9 |
signature BASIC_THM = |
10 |
sig |
|
11 |
include BASIC_THM |
|
12 |
structure Ctermtab: TABLE |
|
13 |
structure Thmtab: TABLE |
|
14 |
val aconvc: cterm * cterm -> bool |
|
45375
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
43780
diff
changeset
|
15 |
type attribute = Context.generic * thm -> Context.generic option * thm option |
32842 | 16 |
end; |
17 |
||
22362
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
18 |
signature THM = |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
19 |
sig |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
20 |
include THM |
32842 | 21 |
structure Ctermtab: TABLE |
22 |
structure Thmtab: TABLE |
|
24948 | 23 |
val aconvc: cterm * cterm -> bool |
23491 | 24 |
val add_cterm_frees: cterm -> cterm list -> cterm list |
32198 | 25 |
val all_name: string * cterm -> cterm -> cterm |
26 |
val all: cterm -> cterm -> cterm |
|
22907 | 27 |
val mk_binop: cterm -> cterm -> cterm -> cterm |
28 |
val dest_binop: cterm -> cterm * cterm |
|
29 |
val dest_implies: cterm -> cterm * cterm |
|
30 |
val dest_equals: cterm -> cterm * cterm |
|
31 |
val dest_equals_lhs: cterm -> cterm |
|
32 |
val dest_equals_rhs: cterm -> cterm |
|
33 |
val lhs_of: thm -> cterm |
|
34 |
val rhs_of: thm -> cterm |
|
22362
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
35 |
val thm_ord: thm * thm -> order |
32842 | 36 |
val cterm_cache: (cterm -> 'a) -> cterm -> 'a |
37 |
val thm_cache: (thm -> 'a) -> thm -> 'a |
|
23599 | 38 |
val is_reflexive: thm -> bool |
22362
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
39 |
val eq_thm: thm * thm -> bool |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
40 |
val eq_thm_prop: thm * thm -> bool |
52683
fb028440473e
more official Thm.eq_thm_strict, without demanding ML equality type;
wenzelm
parents:
51316
diff
changeset
|
41 |
val eq_thm_strict: thm * thm -> bool |
22362
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
42 |
val equiv_thm: thm * thm -> bool |
31944 | 43 |
val class_triv: theory -> class -> thm |
44 |
val of_sort: ctyp * sort -> thm list |
|
28621
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
wenzelm
parents:
28116
diff
changeset
|
45 |
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
|
46 |
val is_dummy: thm -> bool |
22695
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
47 |
val plain_prop_of: thm -> term |
24048
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents:
23599
diff
changeset
|
48 |
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
|
49 |
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
|
50 |
val merge_thms: thm list * thm list -> thm list |
33453 | 51 |
val full_rules: thm Item_Net.T |
30560 | 52 |
val intro_rules: thm Item_Net.T |
53 |
val elim_rules: thm Item_Net.T |
|
54984
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
wenzelm
parents:
53206
diff
changeset
|
54 |
val declare_hyps: cterm -> Proof.context -> Proof.context |
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
wenzelm
parents:
53206
diff
changeset
|
55 |
val assume_hyps: cterm -> Proof.context -> thm * Proof.context |
54993
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
wenzelm
parents:
54984
diff
changeset
|
56 |
val unchecked_hyps: Proof.context -> Proof.context |
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
wenzelm
parents:
54984
diff
changeset
|
57 |
val restore_hyps: Proof.context -> Proof.context -> Proof.context |
55633 | 58 |
val undeclared_hyps: Context.generic -> thm -> term list |
54993
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
wenzelm
parents:
54984
diff
changeset
|
59 |
val check_hyps: Context.generic -> thm -> thm |
27866
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
60 |
val elim_implies: thm -> thm -> thm |
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
61 |
val forall_elim_var: int -> thm -> thm |
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
62 |
val forall_elim_vars: int -> thm -> thm |
32279 | 63 |
val certify_inst: theory -> |
64 |
((indexname * sort) * typ) list * ((indexname * typ) * term) list -> |
|
65 |
(ctyp * ctyp) list * (cterm * cterm) list |
|
66 |
val certify_instantiate: |
|
67 |
((indexname * sort) * typ) list * ((indexname * typ) * term) list -> thm -> thm |
|
35985
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
wenzelm
parents:
35857
diff
changeset
|
68 |
val forall_intr_frees: thm -> thm |
35845
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents:
35715
diff
changeset
|
69 |
val unvarify_global: thm -> thm |
27866
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
70 |
val close_derivation: thm -> thm |
42375
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
40238
diff
changeset
|
71 |
val add_axiom: Proof.context -> binding * term -> theory -> (string * thm) * theory |
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
40238
diff
changeset
|
72 |
val add_axiom_global: binding * term -> theory -> (string * thm) * theory |
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
40238
diff
changeset
|
73 |
val add_def: Proof.context -> bool -> bool -> binding * term -> theory -> (string * thm) * theory |
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
40238
diff
changeset
|
74 |
val add_def_global: bool -> bool -> binding * term -> theory -> (string * thm) * theory |
45375
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
43780
diff
changeset
|
75 |
type attribute = Context.generic * thm -> Context.generic option * thm option |
30210 | 76 |
type binding = binding * attribute list |
77 |
val empty_binding: binding |
|
27866
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
78 |
val rule_attribute: (Context.generic -> thm -> thm) -> attribute |
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
79 |
val declaration_attribute: (thm -> Context.generic -> Context.generic) -> attribute |
45375
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
43780
diff
changeset
|
80 |
val mixed_attribute: (Context.generic * thm -> Context.generic * thm) -> attribute |
46775
6287653e63ec
canonical argument order for attribute application;
wenzelm
parents:
46497
diff
changeset
|
81 |
val apply_attribute: attribute -> thm -> Context.generic -> thm * Context.generic |
45375
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
43780
diff
changeset
|
82 |
val attribute_declaration: attribute -> thm -> Context.generic -> Context.generic |
46775
6287653e63ec
canonical argument order for attribute application;
wenzelm
parents:
46497
diff
changeset
|
83 |
val theory_attributes: attribute list -> thm -> theory -> thm * theory |
6287653e63ec
canonical argument order for attribute application;
wenzelm
parents:
46497
diff
changeset
|
84 |
val proof_attributes: attribute list -> thm -> Proof.context -> thm * Proof.context |
27866
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
85 |
val no_attributes: 'a -> 'a * 'b list |
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
86 |
val simple_fact: 'a -> ('a * 'b list) list |
46830 | 87 |
val tag_rule: string * string -> thm -> thm |
27866
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
88 |
val untag_rule: string -> thm -> thm |
46830 | 89 |
val tag: string * string -> attribute |
27866
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
90 |
val untag: string -> attribute |
30342 | 91 |
val def_name: string -> string |
92 |
val def_name_optional: string -> string -> string |
|
35238 | 93 |
val def_binding: Binding.binding -> Binding.binding |
30433
ce5138c92ca7
added def_binding_optional -- robust version of def_name_optional for bindings;
wenzelm
parents:
30342
diff
changeset
|
94 |
val def_binding_optional: Binding.binding -> Binding.binding -> Binding.binding |
27866
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
95 |
val has_name_hint: thm -> bool |
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
96 |
val get_name_hint: thm -> string |
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
97 |
val put_name_hint: string -> thm -> thm |
22362
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
98 |
val theoremK: string |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
99 |
val lemmaK: string |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
100 |
val corollaryK: string |
42473 | 101 |
val legacy_get_kind: thm -> string |
27866
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
102 |
val kind_rule: string -> thm -> thm |
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
103 |
val kind: string -> attribute |
49062
7e31dfd99ce7
discontinued complicated/unreliable notion of recent proofs within context;
wenzelm
parents:
49058
diff
changeset
|
104 |
val register_proofs: thm list -> theory -> theory |
49011
9c68e43502ce
some support for registering forked proofs within Proof.state, using its bottom context;
wenzelm
parents:
49010
diff
changeset
|
105 |
val join_theory_proofs: theory -> unit |
22362
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
106 |
end; |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
107 |
|
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
108 |
structure Thm: THM = |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
109 |
struct |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
110 |
|
22695
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
111 |
(** basic operations **) |
22362
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
112 |
|
23491 | 113 |
(* collecting cterms *) |
114 |
||
115 |
val op aconvc = op aconv o pairself Thm.term_of; |
|
116 |
||
117 |
fun add_cterm_frees ct = |
|
118 |
let |
|
119 |
val cert = Thm.cterm_of (Thm.theory_of_cterm ct); |
|
120 |
val t = Thm.term_of ct; |
|
121 |
in Term.fold_aterms (fn v as Free _ => insert (op aconvc) (cert v) | _ => I) t end; |
|
122 |
||
123 |
||
22907 | 124 |
(* cterm constructors and destructors *) |
125 |
||
32198 | 126 |
fun all_name (x, t) A = |
127 |
let |
|
128 |
val cert = Thm.cterm_of (Thm.theory_of_cterm t); |
|
129 |
val T = #T (Thm.rep_cterm t); |
|
56245 | 130 |
in Thm.apply (cert (Const ("Pure.all", (T --> propT) --> propT))) (Thm.lambda_name (x, t) A) end; |
32198 | 131 |
|
132 |
fun all t A = all_name ("", t) A; |
|
133 |
||
46497
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
wenzelm
parents:
45382
diff
changeset
|
134 |
fun mk_binop c a b = Thm.apply (Thm.apply c a) b; |
22907 | 135 |
fun dest_binop ct = (Thm.dest_arg1 ct, Thm.dest_arg ct); |
136 |
||
137 |
fun dest_implies ct = |
|
138 |
(case Thm.term_of ct of |
|
56245 | 139 |
Const ("Pure.imp", _) $ _ $ _ => dest_binop ct |
22907 | 140 |
| _ => raise TERM ("dest_implies", [Thm.term_of ct])); |
141 |
||
142 |
fun dest_equals ct = |
|
143 |
(case Thm.term_of ct of |
|
56245 | 144 |
Const ("Pure.eq", _) $ _ $ _ => dest_binop ct |
22907 | 145 |
| _ => raise TERM ("dest_equals", [Thm.term_of ct])); |
146 |
||
147 |
fun dest_equals_lhs ct = |
|
148 |
(case Thm.term_of ct of |
|
56245 | 149 |
Const ("Pure.eq", _) $ _ $ _ => Thm.dest_arg1 ct |
22907 | 150 |
| _ => raise TERM ("dest_equals_lhs", [Thm.term_of ct])); |
151 |
||
152 |
fun dest_equals_rhs ct = |
|
153 |
(case Thm.term_of ct of |
|
56245 | 154 |
Const ("Pure.eq", _) $ _ $ _ => Thm.dest_arg ct |
22907 | 155 |
| _ => raise TERM ("dest_equals_rhs", [Thm.term_of ct])); |
156 |
||
157 |
val lhs_of = dest_equals_lhs o Thm.cprop_of; |
|
158 |
val rhs_of = dest_equals_rhs o Thm.cprop_of; |
|
159 |
||
160 |
||
161 |
(* thm order: ignores theory context! *) |
|
22682 | 162 |
|
22362
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
163 |
fun thm_ord (th1, th2) = |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
164 |
let |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
165 |
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
|
166 |
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
|
167 |
in |
35408 | 168 |
(case Term_Ord.fast_term_ord (prop1, prop2) of |
22362
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
169 |
EQUAL => |
35408 | 170 |
(case list_ord (prod_ord Term_Ord.fast_term_ord Term_Ord.fast_term_ord) (tpairs1, tpairs2) of |
22362
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
171 |
EQUAL => |
35408 | 172 |
(case list_ord Term_Ord.fast_term_ord (hyps1, hyps2) of |
173 |
EQUAL => list_ord Term_Ord.sort_ord (shyps1, shyps2) |
|
22362
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
174 |
| ord => ord) |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
175 |
| ord => ord) |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
176 |
| ord => ord) |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
177 |
end; |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
178 |
|
22682 | 179 |
|
32842 | 180 |
(* tables and caches *) |
181 |
||
35408 | 182 |
structure Ctermtab = Table(type key = cterm val ord = Term_Ord.fast_term_ord o pairself Thm.term_of); |
32842 | 183 |
structure Thmtab = Table(type key = thm val ord = thm_ord); |
184 |
||
185 |
fun cterm_cache f = Cache.create Ctermtab.empty Ctermtab.lookup Ctermtab.update f; |
|
186 |
fun thm_cache f = Cache.create Thmtab.empty Thmtab.lookup Thmtab.update f; |
|
187 |
||
188 |
||
22682 | 189 |
(* equality *) |
190 |
||
23599 | 191 |
fun is_reflexive th = op aconv (Logic.dest_equals (Thm.prop_of th)) |
192 |
handle TERM _ => false; |
|
193 |
||
55547
384bfd19ee61
subtle change of semantics of Thm.eq_thm, e.g. relevant for merge of src/HOL/Tools/Predicate_Compile/core_data.ML (cf. HOL-IMP);
wenzelm
parents:
54996
diff
changeset
|
194 |
val eq_thm = is_equal o thm_ord; |
22362
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
195 |
|
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
196 |
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
|
197 |
|
52683
fb028440473e
more official Thm.eq_thm_strict, without demanding ML equality type;
wenzelm
parents:
51316
diff
changeset
|
198 |
fun eq_thm_strict ths = |
55547
384bfd19ee61
subtle change of semantics of Thm.eq_thm, e.g. relevant for merge of src/HOL/Tools/Predicate_Compile/core_data.ML (cf. HOL-IMP);
wenzelm
parents:
54996
diff
changeset
|
199 |
eq_thm ths andalso |
384bfd19ee61
subtle change of semantics of Thm.eq_thm, e.g. relevant for merge of src/HOL/Tools/Predicate_Compile/core_data.ML (cf. HOL-IMP);
wenzelm
parents:
54996
diff
changeset
|
200 |
let val (rep1, rep2) = pairself Thm.rep_thm ths in |
384bfd19ee61
subtle change of semantics of Thm.eq_thm, e.g. relevant for merge of src/HOL/Tools/Predicate_Compile/core_data.ML (cf. HOL-IMP);
wenzelm
parents:
54996
diff
changeset
|
201 |
Theory.eq_thy (#thy rep1, #thy rep2) andalso |
384bfd19ee61
subtle change of semantics of Thm.eq_thm, e.g. relevant for merge of src/HOL/Tools/Predicate_Compile/core_data.ML (cf. HOL-IMP);
wenzelm
parents:
54996
diff
changeset
|
202 |
#maxidx rep1 = #maxidx rep2 andalso |
384bfd19ee61
subtle change of semantics of Thm.eq_thm, e.g. relevant for merge of src/HOL/Tools/Predicate_Compile/core_data.ML (cf. HOL-IMP);
wenzelm
parents:
54996
diff
changeset
|
203 |
#tags rep1 = #tags rep2 |
384bfd19ee61
subtle change of semantics of Thm.eq_thm, e.g. relevant for merge of src/HOL/Tools/Predicate_Compile/core_data.ML (cf. HOL-IMP);
wenzelm
parents:
54996
diff
changeset
|
204 |
end; |
52683
fb028440473e
more official Thm.eq_thm_strict, without demanding ML equality type;
wenzelm
parents:
51316
diff
changeset
|
205 |
|
22682 | 206 |
|
207 |
(* pattern equivalence *) |
|
208 |
||
22362
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
209 |
fun equiv_thm ths = |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
210 |
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
|
211 |
|
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
212 |
|
31904
a86896359ca4
renamed Drule.sort_triv to Thm.sort_triv (cf. more_thm.ML);
wenzelm
parents:
31177
diff
changeset
|
213 |
(* type classes and sorts *) |
a86896359ca4
renamed Drule.sort_triv to Thm.sort_triv (cf. more_thm.ML);
wenzelm
parents:
31177
diff
changeset
|
214 |
|
31944 | 215 |
fun class_triv thy c = |
216 |
Thm.of_class (Thm.ctyp_of thy (TVar ((Name.aT, 0), [c])), c); |
|
217 |
||
218 |
fun of_sort (T, S) = map (fn c => Thm.of_class (T, c)) S; |
|
28621
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
wenzelm
parents:
28116
diff
changeset
|
219 |
|
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
wenzelm
parents:
28116
diff
changeset
|
220 |
fun check_shyps sorts raw_th = |
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
wenzelm
parents:
28116
diff
changeset
|
221 |
let |
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
wenzelm
parents:
28116
diff
changeset
|
222 |
val th = Thm.strip_shyps raw_th; |
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
wenzelm
parents:
28116
diff
changeset
|
223 |
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
|
224 |
val pending = Sorts.subtract sorts (Thm.extra_shyps th); |
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
wenzelm
parents:
28116
diff
changeset
|
225 |
in |
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
wenzelm
parents:
28116
diff
changeset
|
226 |
if null pending then th |
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
wenzelm
parents:
28116
diff
changeset
|
227 |
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
|
228 |
Pretty.brk 1 :: Pretty.commas (map prt_sort pending)))) |
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
wenzelm
parents:
28116
diff
changeset
|
229 |
end; |
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
wenzelm
parents:
28116
diff
changeset
|
230 |
|
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
wenzelm
parents:
28116
diff
changeset
|
231 |
|
22695
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
232 |
(* misc operations *) |
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
233 |
|
24048
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents:
23599
diff
changeset
|
234 |
fun is_dummy thm = |
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents:
23599
diff
changeset
|
235 |
(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
|
236 |
NONE => false |
58001
934d85f14d1d
more general dummy: may contain "parked arguments", for example;
wenzelm
parents:
56245
diff
changeset
|
237 |
| SOME t => Term.is_dummy_pattern (Term.head_of t)); |
24048
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents:
23599
diff
changeset
|
238 |
|
22695
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
239 |
fun plain_prop_of raw_thm = |
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
240 |
let |
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
241 |
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
|
242 |
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
|
243 |
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
|
244 |
in |
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
245 |
if not (null hyps) then |
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
246 |
err "theorem may not contain hypotheses" |
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
247 |
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
|
248 |
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
|
249 |
else if not (null tpairs) then |
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
250 |
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
|
251 |
else prop |
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
252 |
end; |
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
253 |
|
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
254 |
|
30564 | 255 |
(* collections of theorems in canonical order *) |
24048
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents:
23599
diff
changeset
|
256 |
|
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents:
23599
diff
changeset
|
257 |
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
|
258 |
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
|
259 |
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
|
260 |
|
33453 | 261 |
val full_rules = Item_Net.init eq_thm_prop (single o Thm.full_prop_of); |
33373 | 262 |
val intro_rules = Item_Net.init eq_thm_prop (single o Thm.concl_of); |
263 |
val elim_rules = Item_Net.init eq_thm_prop (single o Thm.major_prem_of); |
|
30560 | 264 |
|
265 |
||
22682 | 266 |
|
54984
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
wenzelm
parents:
53206
diff
changeset
|
267 |
(** declared hyps **) |
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
wenzelm
parents:
53206
diff
changeset
|
268 |
|
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
wenzelm
parents:
53206
diff
changeset
|
269 |
structure Hyps = Proof_Data |
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
wenzelm
parents:
53206
diff
changeset
|
270 |
( |
54993
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
wenzelm
parents:
54984
diff
changeset
|
271 |
type T = Termtab.set * bool; |
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
wenzelm
parents:
54984
diff
changeset
|
272 |
fun init _ : T = (Termtab.empty, true); |
54984
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
wenzelm
parents:
53206
diff
changeset
|
273 |
); |
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
wenzelm
parents:
53206
diff
changeset
|
274 |
|
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
wenzelm
parents:
53206
diff
changeset
|
275 |
fun declare_hyps ct ctxt = |
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
wenzelm
parents:
53206
diff
changeset
|
276 |
if Theory.subthy (theory_of_cterm ct, Proof_Context.theory_of ctxt) then |
54993
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
wenzelm
parents:
54984
diff
changeset
|
277 |
(Hyps.map o apfst) (Termtab.update (term_of ct, ())) ctxt |
54984
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
wenzelm
parents:
53206
diff
changeset
|
278 |
else raise CTERM ("assume_hyps: bad background theory", [ct]); |
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
wenzelm
parents:
53206
diff
changeset
|
279 |
|
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
wenzelm
parents:
53206
diff
changeset
|
280 |
fun assume_hyps ct ctxt = (Thm.assume ct, declare_hyps ct ctxt); |
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
wenzelm
parents:
53206
diff
changeset
|
281 |
|
54993
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
wenzelm
parents:
54984
diff
changeset
|
282 |
val unchecked_hyps = (Hyps.map o apsnd) (K false); |
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
wenzelm
parents:
54984
diff
changeset
|
283 |
fun restore_hyps ctxt = (Hyps.map o apsnd) (K (#2 (Hyps.get ctxt))); |
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
wenzelm
parents:
54984
diff
changeset
|
284 |
|
55633 | 285 |
fun undeclared_hyps context th = |
286 |
Thm.hyps_of th |
|
287 |
|> filter_out |
|
288 |
(case context of |
|
289 |
Context.Theory _ => K false |
|
290 |
| Context.Proof ctxt => |
|
291 |
(case Hyps.get ctxt of |
|
292 |
(_, false) => K true |
|
293 |
| (hyps, _) => Termtab.defined hyps)); |
|
294 |
||
54993
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
wenzelm
parents:
54984
diff
changeset
|
295 |
fun check_hyps context th = |
55633 | 296 |
(case undeclared_hyps context th of |
297 |
[] => th |
|
298 |
| undeclared => |
|
54993
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
wenzelm
parents:
54984
diff
changeset
|
299 |
let |
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
wenzelm
parents:
54984
diff
changeset
|
300 |
val ctxt = Context.cases Syntax.init_pretty_global I context; |
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
wenzelm
parents:
54984
diff
changeset
|
301 |
in |
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
wenzelm
parents:
54984
diff
changeset
|
302 |
error (Pretty.string_of (Pretty.big_list "Undeclared hyps:" |
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
wenzelm
parents:
54984
diff
changeset
|
303 |
(map (Pretty.item o single o Syntax.pretty_term ctxt) undeclared))) |
55633 | 304 |
end); |
54984
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
wenzelm
parents:
53206
diff
changeset
|
305 |
|
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
wenzelm
parents:
53206
diff
changeset
|
306 |
|
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
wenzelm
parents:
53206
diff
changeset
|
307 |
|
24980
16a74cfca971
added elim_implies (more convenient argument order);
wenzelm
parents:
24948
diff
changeset
|
308 |
(** basic derived rules **) |
16a74cfca971
added elim_implies (more convenient argument order);
wenzelm
parents:
24948
diff
changeset
|
309 |
|
16a74cfca971
added elim_implies (more convenient argument order);
wenzelm
parents:
24948
diff
changeset
|
310 |
(*Elimination of implication |
16a74cfca971
added elim_implies (more convenient argument order);
wenzelm
parents:
24948
diff
changeset
|
311 |
A A ==> B |
16a74cfca971
added elim_implies (more convenient argument order);
wenzelm
parents:
24948
diff
changeset
|
312 |
------------ |
16a74cfca971
added elim_implies (more convenient argument order);
wenzelm
parents:
24948
diff
changeset
|
313 |
B |
16a74cfca971
added elim_implies (more convenient argument order);
wenzelm
parents:
24948
diff
changeset
|
314 |
*) |
16a74cfca971
added elim_implies (more convenient argument order);
wenzelm
parents:
24948
diff
changeset
|
315 |
fun elim_implies thA thAB = Thm.implies_elim thAB thA; |
16a74cfca971
added elim_implies (more convenient argument order);
wenzelm
parents:
24948
diff
changeset
|
316 |
|
26653 | 317 |
|
318 |
(* forall_elim_var(s) *) |
|
319 |
||
320 |
local |
|
321 |
||
322 |
fun forall_elim_vars_aux strip_vars i th = |
|
323 |
let |
|
324 |
val thy = Thm.theory_of_thm th; |
|
325 |
val {tpairs, prop, ...} = Thm.rep_thm th; |
|
326 |
val add_used = Term.fold_aterms |
|
327 |
(fn Var ((x, j), _) => if i = j then insert (op =) x else I | _ => I); |
|
328 |
val used = fold (fn (t, u) => add_used t o add_used u) tpairs (add_used prop []); |
|
329 |
val vars = strip_vars prop; |
|
330 |
val cvars = (Name.variant_list used (map #1 vars), vars) |
|
331 |
|> ListPair.map (fn (x, (_, T)) => Thm.cterm_of thy (Var ((x, i), T))); |
|
332 |
in fold Thm.forall_elim cvars th end; |
|
333 |
||
334 |
in |
|
335 |
||
336 |
val forall_elim_vars = forall_elim_vars_aux Term.strip_all_vars; |
|
337 |
||
33697 | 338 |
fun forall_elim_var i th = |
339 |
forall_elim_vars_aux |
|
56245 | 340 |
(fn Const ("Pure.all", _) $ Abs (a, T, _) => [(a, T)] |
33697 | 341 |
| _ => raise THM ("forall_elim_vars", i, [th])) i th; |
26653 | 342 |
|
343 |
end; |
|
344 |
||
345 |
||
32279 | 346 |
(* certify_instantiate *) |
347 |
||
348 |
fun certify_inst thy (instT, inst) = |
|
349 |
(map (fn (v, T) => (Thm.ctyp_of thy (TVar v), Thm.ctyp_of thy T)) instT, |
|
350 |
map (fn (v, t) => (Thm.cterm_of thy (Var v), Thm.cterm_of thy t)) inst); |
|
351 |
||
352 |
fun certify_instantiate insts th = |
|
353 |
Thm.instantiate (certify_inst (Thm.theory_of_thm th) insts) th; |
|
354 |
||
355 |
||
35985
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
wenzelm
parents:
35857
diff
changeset
|
356 |
(* forall_intr_frees: generalization over all suitable Free variables *) |
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
wenzelm
parents:
35857
diff
changeset
|
357 |
|
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
wenzelm
parents:
35857
diff
changeset
|
358 |
fun forall_intr_frees th = |
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
wenzelm
parents:
35857
diff
changeset
|
359 |
let |
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
wenzelm
parents:
35857
diff
changeset
|
360 |
val thy = Thm.theory_of_thm th; |
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
wenzelm
parents:
35857
diff
changeset
|
361 |
val {prop, hyps, tpairs, ...} = Thm.rep_thm th; |
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
wenzelm
parents:
35857
diff
changeset
|
362 |
val fixed = fold Term.add_frees (Thm.terms_of_tpairs tpairs @ hyps) []; |
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
wenzelm
parents:
35857
diff
changeset
|
363 |
val frees = Term.fold_aterms (fn Free v => |
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
wenzelm
parents:
35857
diff
changeset
|
364 |
if member (op =) fixed v then I else insert (op =) v | _ => I) prop []; |
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
wenzelm
parents:
35857
diff
changeset
|
365 |
in fold (Thm.forall_intr o Thm.cterm_of thy o Free) frees th end; |
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
wenzelm
parents:
35857
diff
changeset
|
366 |
|
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
wenzelm
parents:
35857
diff
changeset
|
367 |
|
35845
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents:
35715
diff
changeset
|
368 |
(* unvarify_global: global schematic variables *) |
26653 | 369 |
|
35845
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents:
35715
diff
changeset
|
370 |
fun unvarify_global th = |
24980
16a74cfca971
added elim_implies (more convenient argument order);
wenzelm
parents:
24948
diff
changeset
|
371 |
let |
16a74cfca971
added elim_implies (more convenient argument order);
wenzelm
parents:
24948
diff
changeset
|
372 |
val prop = Thm.full_prop_of th; |
35845
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents:
35715
diff
changeset
|
373 |
val _ = map Logic.unvarify_global (prop :: Thm.hyps_of th) |
24980
16a74cfca971
added elim_implies (more convenient argument order);
wenzelm
parents:
24948
diff
changeset
|
374 |
handle TERM (msg, _) => raise THM (msg, 0, [th]); |
16a74cfca971
added elim_implies (more convenient argument order);
wenzelm
parents:
24948
diff
changeset
|
375 |
|
32279 | 376 |
val instT = rev (Term.add_tvars prop []) |> map (fn v as ((a, _), S) => (v, TFree (a, S))); |
24980
16a74cfca971
added elim_implies (more convenient argument order);
wenzelm
parents:
24948
diff
changeset
|
377 |
val inst = rev (Term.add_vars prop []) |> map (fn ((a, i), T) => |
32279 | 378 |
let val T' = Term_Subst.instantiateT instT T |
379 |
in (((a, i), T'), Free ((a, T'))) end); |
|
380 |
in certify_instantiate (instT, inst) th end; |
|
24980
16a74cfca971
added elim_implies (more convenient argument order);
wenzelm
parents:
24948
diff
changeset
|
381 |
|
26653 | 382 |
|
383 |
(* close_derivation *) |
|
384 |
||
26628
63306cb94313
replaced Drule.close_derivation/Goal.close_result by Thm.close_derivation (removed obsolete compression);
wenzelm
parents:
25518
diff
changeset
|
385 |
fun close_derivation thm = |
36744
6e1f3d609a68
renamed Thm.get_name -> Thm.derivation_name and Thm.put_name -> Thm.name_derivation, to emphasize the true nature of these operations;
wenzelm
parents:
36106
diff
changeset
|
386 |
if Thm.derivation_name thm = "" then Thm.name_derivation "" thm |
26628
63306cb94313
replaced Drule.close_derivation/Goal.close_result by Thm.close_derivation (removed obsolete compression);
wenzelm
parents:
25518
diff
changeset
|
387 |
else thm; |
63306cb94313
replaced Drule.close_derivation/Goal.close_result by Thm.close_derivation (removed obsolete compression);
wenzelm
parents:
25518
diff
changeset
|
388 |
|
24980
16a74cfca971
added elim_implies (more convenient argument order);
wenzelm
parents:
24948
diff
changeset
|
389 |
|
16a74cfca971
added elim_implies (more convenient argument order);
wenzelm
parents:
24948
diff
changeset
|
390 |
|
16a74cfca971
added elim_implies (more convenient argument order);
wenzelm
parents:
24948
diff
changeset
|
391 |
(** specification primitives **) |
16a74cfca971
added elim_implies (more convenient argument order);
wenzelm
parents:
24948
diff
changeset
|
392 |
|
30342 | 393 |
(* rules *) |
394 |
||
35855
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
wenzelm
parents:
35853
diff
changeset
|
395 |
fun stripped_sorts thy t = |
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
wenzelm
parents:
35853
diff
changeset
|
396 |
let |
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
wenzelm
parents:
35853
diff
changeset
|
397 |
val tfrees = rev (map TFree (Term.add_tfrees t [])); |
43329
84472e198515
tuned signature: Name.invent and Name.invent_names;
wenzelm
parents:
42473
diff
changeset
|
398 |
val tfrees' = map (fn a => TFree (a, [])) (Name.invent Name.context Name.aT (length tfrees)); |
35855
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
wenzelm
parents:
35853
diff
changeset
|
399 |
val strip = tfrees ~~ tfrees'; |
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
wenzelm
parents:
35853
diff
changeset
|
400 |
val recover = map (pairself (Thm.ctyp_of thy o Logic.varifyT_global) o swap) strip; |
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
wenzelm
parents:
35853
diff
changeset
|
401 |
val t' = Term.map_types (Term.map_atyps (perhaps (AList.lookup (op =) strip))) t; |
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
wenzelm
parents:
35853
diff
changeset
|
402 |
in (strip, recover, t') end; |
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
wenzelm
parents:
35853
diff
changeset
|
403 |
|
42375
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
40238
diff
changeset
|
404 |
fun add_axiom ctxt (b, prop) thy = |
24980
16a74cfca971
added elim_implies (more convenient argument order);
wenzelm
parents:
24948
diff
changeset
|
405 |
let |
42375
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
40238
diff
changeset
|
406 |
val _ = Sign.no_vars ctxt prop; |
35855
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
wenzelm
parents:
35853
diff
changeset
|
407 |
val (strip, recover, prop') = stripped_sorts thy prop; |
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
wenzelm
parents:
35853
diff
changeset
|
408 |
val constraints = map (fn (TFree (_, S), T) => (T, S)) strip; |
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
wenzelm
parents:
35853
diff
changeset
|
409 |
val of_sorts = maps (fn (T as TFree (_, S), _) => of_sort (Thm.ctyp_of thy T, S)) strip; |
36106
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
wenzelm
parents:
35988
diff
changeset
|
410 |
|
42375
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
40238
diff
changeset
|
411 |
val thy' = thy |
51316
dfe469293eb4
discontinued empty name bindings in 'axiomatization';
wenzelm
parents:
49062
diff
changeset
|
412 |
|> Theory.add_axiom ctxt (b, Logic.list_implies (maps Logic.mk_of_sort constraints, prop')); |
dfe469293eb4
discontinued empty name bindings in 'axiomatization';
wenzelm
parents:
49062
diff
changeset
|
413 |
val axm_name = Sign.full_name thy' b; |
36106
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
wenzelm
parents:
35988
diff
changeset
|
414 |
val axm' = Thm.axiom thy' axm_name; |
35988
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
wenzelm
parents:
35985
diff
changeset
|
415 |
val thm = |
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
wenzelm
parents:
35985
diff
changeset
|
416 |
Thm.instantiate (recover, []) axm' |
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
wenzelm
parents:
35985
diff
changeset
|
417 |
|> unvarify_global |
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
wenzelm
parents:
35985
diff
changeset
|
418 |
|> fold elim_implies of_sorts; |
36106
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
wenzelm
parents:
35988
diff
changeset
|
419 |
in ((axm_name, thm), thy') end; |
24980
16a74cfca971
added elim_implies (more convenient argument order);
wenzelm
parents:
24948
diff
changeset
|
420 |
|
42375
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
40238
diff
changeset
|
421 |
fun add_axiom_global arg thy = add_axiom (Syntax.init_pretty_global thy) arg thy; |
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
40238
diff
changeset
|
422 |
|
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
40238
diff
changeset
|
423 |
fun add_def ctxt unchecked overloaded (b, prop) thy = |
24980
16a74cfca971
added elim_implies (more convenient argument order);
wenzelm
parents:
24948
diff
changeset
|
424 |
let |
42375
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
40238
diff
changeset
|
425 |
val _ = Sign.no_vars ctxt prop; |
35988
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
wenzelm
parents:
35985
diff
changeset
|
426 |
val prems = map (Thm.cterm_of thy) (Logic.strip_imp_prems prop); |
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
wenzelm
parents:
35985
diff
changeset
|
427 |
val (_, recover, concl') = stripped_sorts thy (Logic.strip_imp_concl prop); |
36106
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
wenzelm
parents:
35988
diff
changeset
|
428 |
|
42375
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
40238
diff
changeset
|
429 |
val thy' = Theory.add_def ctxt unchecked overloaded (b, concl') thy; |
36106
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
wenzelm
parents:
35988
diff
changeset
|
430 |
val axm_name = Sign.full_name thy' b; |
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
wenzelm
parents:
35988
diff
changeset
|
431 |
val axm' = Thm.axiom thy' axm_name; |
35988
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
wenzelm
parents:
35985
diff
changeset
|
432 |
val thm = |
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
wenzelm
parents:
35985
diff
changeset
|
433 |
Thm.instantiate (recover, []) axm' |
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
wenzelm
parents:
35985
diff
changeset
|
434 |
|> unvarify_global |
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
wenzelm
parents:
35985
diff
changeset
|
435 |
|> fold_rev Thm.implies_intr prems; |
36106
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
wenzelm
parents:
35988
diff
changeset
|
436 |
in ((axm_name, thm), thy') end; |
24980
16a74cfca971
added elim_implies (more convenient argument order);
wenzelm
parents:
24948
diff
changeset
|
437 |
|
42375
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
40238
diff
changeset
|
438 |
fun add_def_global unchecked overloaded arg thy = |
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
40238
diff
changeset
|
439 |
add_def (Syntax.init_pretty_global thy) unchecked overloaded arg thy; |
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
40238
diff
changeset
|
440 |
|
27866
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
441 |
|
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
442 |
|
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
443 |
(** attributes **) |
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
444 |
|
40238
edcdecd55655
type attribute is derived concept outside the kernel;
wenzelm
parents:
39133
diff
changeset
|
445 |
(*attributes subsume any kind of rules or context modifiers*) |
45375
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
43780
diff
changeset
|
446 |
type attribute = Context.generic * thm -> Context.generic option * thm option; |
40238
edcdecd55655
type attribute is derived concept outside the kernel;
wenzelm
parents:
39133
diff
changeset
|
447 |
|
30210 | 448 |
type binding = binding * attribute list; |
449 |
val empty_binding: binding = (Binding.empty, []); |
|
450 |
||
45375
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
43780
diff
changeset
|
451 |
fun rule_attribute f (x, th) = (NONE, SOME (f x th)); |
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
43780
diff
changeset
|
452 |
fun declaration_attribute f (x, th) = (SOME (f th x), NONE); |
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
43780
diff
changeset
|
453 |
fun mixed_attribute f (x, th) = let val (x', th') = f (x, th) in (SOME x', SOME th') end; |
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
43780
diff
changeset
|
454 |
|
46775
6287653e63ec
canonical argument order for attribute application;
wenzelm
parents:
46497
diff
changeset
|
455 |
fun apply_attribute (att: attribute) th x = |
54996 | 456 |
let val (x', th') = att (x, check_hyps x (Thm.transfer (Context.theory_of x) th)) |
46775
6287653e63ec
canonical argument order for attribute application;
wenzelm
parents:
46497
diff
changeset
|
457 |
in (the_default th th', the_default x x') end; |
45375
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
43780
diff
changeset
|
458 |
|
46775
6287653e63ec
canonical argument order for attribute application;
wenzelm
parents:
46497
diff
changeset
|
459 |
fun attribute_declaration att th x = #2 (apply_attribute att th x); |
27866
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
460 |
|
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
461 |
fun apply_attributes mk dest = |
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
462 |
let |
46775
6287653e63ec
canonical argument order for attribute application;
wenzelm
parents:
46497
diff
changeset
|
463 |
fun app [] th x = (th, x) |
6287653e63ec
canonical argument order for attribute application;
wenzelm
parents:
46497
diff
changeset
|
464 |
| app (att :: atts) th x = apply_attribute att th (mk x) ||> dest |-> app atts; |
27866
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
465 |
in app end; |
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
466 |
|
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
467 |
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
|
468 |
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
|
469 |
|
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
470 |
fun no_attributes x = (x, []); |
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
471 |
fun simple_fact x = [(x, [])]; |
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
472 |
|
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
473 |
|
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
474 |
|
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
475 |
(*** theorem tags ***) |
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
476 |
|
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
477 |
(* add / delete tags *) |
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
478 |
|
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
479 |
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
|
480 |
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
|
481 |
|
45375
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
43780
diff
changeset
|
482 |
fun tag tg = rule_attribute (K (tag_rule tg)); |
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
43780
diff
changeset
|
483 |
fun untag s = rule_attribute (K (untag_rule s)); |
27866
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
484 |
|
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
485 |
|
30342 | 486 |
(* def_name *) |
487 |
||
488 |
fun def_name c = c ^ "_def"; |
|
489 |
||
490 |
fun def_name_optional c "" = def_name c |
|
491 |
| def_name_optional _ name = name; |
|
492 |
||
35238 | 493 |
val def_binding = Binding.map_name def_name; |
494 |
||
30433
ce5138c92ca7
added def_binding_optional -- robust version of def_name_optional for bindings;
wenzelm
parents:
30342
diff
changeset
|
495 |
fun def_binding_optional b name = |
35238 | 496 |
if Binding.is_empty name then def_binding b else name; |
30433
ce5138c92ca7
added def_binding_optional -- robust version of def_name_optional for bindings;
wenzelm
parents:
30342
diff
changeset
|
497 |
|
30342 | 498 |
|
27866
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
499 |
(* unofficial theorem names *) |
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
500 |
|
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
501 |
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
|
502 |
|
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
503 |
val has_name_hint = can the_name_hint; |
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
504 |
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
|
505 |
|
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
506 |
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
|
507 |
|
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
508 |
|
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
509 |
(* theorem kinds *) |
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
510 |
|
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
511 |
val theoremK = "theorem"; |
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
512 |
val lemmaK = "lemma"; |
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
513 |
val corollaryK = "corollary"; |
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
514 |
|
42473 | 515 |
fun legacy_get_kind thm = the_default "" (Properties.get (Thm.get_tags thm) Markup.kindN); |
27866
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
516 |
|
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
517 |
fun kind_rule k = tag_rule (Markup.kindN, k) o untag_rule Markup.kindN; |
45375
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
43780
diff
changeset
|
518 |
fun kind k = rule_attribute (K (k <> "" ? kind_rule k)); |
27866
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
519 |
|
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
wenzelm
parents:
27255
diff
changeset
|
520 |
|
49011
9c68e43502ce
some support for registering forked proofs within Proof.state, using its bottom context;
wenzelm
parents:
49010
diff
changeset
|
521 |
(* forked proofs *) |
49010 | 522 |
|
49062
7e31dfd99ce7
discontinued complicated/unreliable notion of recent proofs within context;
wenzelm
parents:
49058
diff
changeset
|
523 |
structure Proofs = Theory_Data |
49010 | 524 |
( |
49062
7e31dfd99ce7
discontinued complicated/unreliable notion of recent proofs within context;
wenzelm
parents:
49058
diff
changeset
|
525 |
type T = thm list; |
7e31dfd99ce7
discontinued complicated/unreliable notion of recent proofs within context;
wenzelm
parents:
49058
diff
changeset
|
526 |
val empty = []; |
49010 | 527 |
fun extend _ = empty; |
528 |
fun merge _ = empty; |
|
529 |
); |
|
530 |
||
49062
7e31dfd99ce7
discontinued complicated/unreliable notion of recent proofs within context;
wenzelm
parents:
49058
diff
changeset
|
531 |
fun register_proofs more_thms = Proofs.map (fn thms => fold cons more_thms thms); |
7e31dfd99ce7
discontinued complicated/unreliable notion of recent proofs within context;
wenzelm
parents:
49058
diff
changeset
|
532 |
val join_theory_proofs = Thm.join_proofs o rev o Proofs.get; |
49010 | 533 |
|
534 |
||
22362
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
535 |
open Thm; |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
536 |
|
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
537 |
end; |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
538 |
|
32842 | 539 |
structure Basic_Thm: BASIC_THM = Thm; |
540 |
open Basic_Thm; |
|
23170 | 541 |