| author | nipkow | 
| Sun, 01 Jan 2012 18:12:11 +0100 | |
| changeset 46067 | a03bf644cb27 | 
| parent 45443 | c8a9a5e577bd | 
| child 46217 | 7b19666f0e3d | 
| permissions | -rw-r--r-- | 
| 250 | 1  | 
(* Title: Pure/thm.ML  | 
2  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
|
| 
29269
 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 
wenzelm 
parents: 
29003 
diff
changeset
 | 
3  | 
Author: Makarius  | 
| 
229
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
225 
diff
changeset
 | 
4  | 
|
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
5  | 
The very core of Isabelle's Meta Logic: certified types and terms,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
6  | 
derivations, theorems, framework rules (including lifting and  | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
7  | 
resolution), oracles.  | 
| 0 | 8  | 
*)  | 
9  | 
||
| 6089 | 10  | 
signature BASIC_THM =  | 
| 1503 | 11  | 
sig  | 
| 1160 | 12  | 
(*certified types*)  | 
| 
387
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
13  | 
type ctyp  | 
| 16656 | 14  | 
val rep_ctyp: ctyp ->  | 
| 
26631
 
d6b6c74a8bcf
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
26424 
diff
changeset
 | 
15  | 
   {thy_ref: theory_ref,
 | 
| 16656 | 16  | 
T: typ,  | 
| 20512 | 17  | 
maxidx: int,  | 
| 39687 | 18  | 
sorts: sort Ord_List.T}  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
19  | 
val theory_of_ctyp: ctyp -> theory  | 
| 
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
20  | 
val typ_of: ctyp -> typ  | 
| 
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
21  | 
val ctyp_of: theory -> typ -> ctyp  | 
| 1160 | 22  | 
|
23  | 
(*certified terms*)  | 
|
24  | 
type cterm  | 
|
| 22584 | 25  | 
exception CTERM of string * cterm list  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
26  | 
val rep_cterm: cterm ->  | 
| 
26631
 
d6b6c74a8bcf
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
26424 
diff
changeset
 | 
27  | 
   {thy_ref: theory_ref,
 | 
| 16656 | 28  | 
t: term,  | 
29  | 
T: typ,  | 
|
30  | 
maxidx: int,  | 
|
| 39687 | 31  | 
sorts: sort Ord_List.T}  | 
| 28354 | 32  | 
val crep_cterm: cterm ->  | 
| 39687 | 33  | 
    {thy_ref: theory_ref, t: term, T: ctyp, maxidx: int, sorts: sort Ord_List.T}
 | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
34  | 
val theory_of_cterm: cterm -> theory  | 
| 
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
35  | 
val term_of: cterm -> term  | 
| 
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
36  | 
val cterm_of: theory -> term -> cterm  | 
| 
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
37  | 
val ctyp_of_term: cterm -> ctyp  | 
| 1160 | 38  | 
|
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
39  | 
(*theorems*)  | 
| 1160 | 40  | 
type thm  | 
| 23601 | 41  | 
type conv = cterm -> thm  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
42  | 
val rep_thm: thm ->  | 
| 
26631
 
d6b6c74a8bcf
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
26424 
diff
changeset
 | 
43  | 
   {thy_ref: theory_ref,
 | 
| 28017 | 44  | 
tags: Properties.T,  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
45  | 
maxidx: int,  | 
| 39687 | 46  | 
shyps: sort Ord_List.T,  | 
47  | 
hyps: term Ord_List.T,  | 
|
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
48  | 
tpairs: (term * term) list,  | 
| 
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
49  | 
prop: term}  | 
| 
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
50  | 
val crep_thm: thm ->  | 
| 
26631
 
d6b6c74a8bcf
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
26424 
diff
changeset
 | 
51  | 
   {thy_ref: theory_ref,
 | 
| 28017 | 52  | 
tags: Properties.T,  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
53  | 
maxidx: int,  | 
| 39687 | 54  | 
shyps: sort Ord_List.T,  | 
55  | 
hyps: cterm Ord_List.T,  | 
|
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
56  | 
tpairs: (cterm * cterm) list,  | 
| 
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
57  | 
prop: cterm}  | 
| 6089 | 58  | 
exception THM of string * int * thm list  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
59  | 
val theory_of_thm: thm -> theory  | 
| 
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
60  | 
val prop_of: thm -> term  | 
| 16656 | 61  | 
val concl_of: thm -> term  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
62  | 
val prems_of: thm -> term list  | 
| 
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
63  | 
val nprems_of: thm -> int  | 
| 
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
64  | 
val cprop_of: thm -> cterm  | 
| 18145 | 65  | 
val cprem_of: thm -> int -> cterm  | 
| 250 | 66  | 
end;  | 
| 0 | 67  | 
|
| 6089 | 68  | 
signature THM =  | 
69  | 
sig  | 
|
70  | 
include BASIC_THM  | 
|
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
71  | 
val dest_ctyp: ctyp -> ctyp list  | 
| 
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
72  | 
val dest_comb: cterm -> cterm * cterm  | 
| 22909 | 73  | 
val dest_fun: cterm -> cterm  | 
| 
20580
 
6fb75df09253
added dest_arg, i.e. a tuned version of #2 o dest_comb;
 
wenzelm 
parents: 
20548 
diff
changeset
 | 
74  | 
val dest_arg: cterm -> cterm  | 
| 22909 | 75  | 
val dest_fun2: cterm -> cterm  | 
76  | 
val dest_arg1: cterm -> cterm  | 
|
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
77  | 
val dest_abs: string option -> cterm -> cterm * cterm  | 
| 
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
78  | 
val capply: cterm -> cterm -> cterm  | 
| 32198 | 79  | 
val cabs_name: string * cterm -> cterm -> cterm  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
80  | 
val cabs: cterm -> cterm -> cterm  | 
| 31945 | 81  | 
val adjust_maxidx_cterm: int -> cterm -> cterm  | 
82  | 
val incr_indexes_cterm: int -> cterm -> cterm  | 
|
83  | 
val match: cterm * cterm -> (ctyp * ctyp) list * (cterm * cterm) list  | 
|
84  | 
val first_order_match: cterm * cterm -> (ctyp * ctyp) list * (cterm * cterm) list  | 
|
| 
31947
 
7daee3bed3af
clarified strip_shyps: proper type witnesses for present sorts;
 
wenzelm 
parents: 
31945 
diff
changeset
 | 
85  | 
val fold_terms: (term -> 'a -> 'a) -> thm -> 'a -> 'a  | 
| 16945 | 86  | 
val terms_of_tpairs: (term * term) list -> term list  | 
| 31945 | 87  | 
val full_prop_of: thm -> term  | 
| 19881 | 88  | 
val maxidx_of: thm -> int  | 
| 19910 | 89  | 
val maxidx_thm: thm -> int -> int  | 
| 19881 | 90  | 
val hyps_of: thm -> term list  | 
| 38709 | 91  | 
val tpairs_of: thm -> (term * term) list  | 
| 31945 | 92  | 
val no_prems: thm -> bool  | 
93  | 
val major_prem_of: thm -> term  | 
|
| 38709 | 94  | 
val transfer: theory -> thm -> thm  | 
95  | 
val weaken: cterm -> thm -> thm  | 
|
96  | 
val weaken_sorts: sort list -> cterm -> cterm  | 
|
97  | 
val extra_shyps: thm -> sort list  | 
|
| 
44333
 
cc53ce50f738
reverted to join_bodies/join_proofs based on fold_body_thms to regain performance (escpecially of HOL-Proofs) -- see also aa9c1e9ef2ce and 4e2abb045eac;
 
wenzelm 
parents: 
44331 
diff
changeset
 | 
98  | 
val proof_bodies_of: thm list -> proof_body list  | 
| 32725 | 99  | 
val proof_body_of: thm -> proof_body  | 
100  | 
val proof_of: thm -> proof  | 
|
| 44331 | 101  | 
val join_proofs: thm list -> unit  | 
| 32725 | 102  | 
  val status_of: thm -> {oracle: bool, unfinished: bool, failed: bool}
 | 
103  | 
val future: thm future -> cterm -> 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: 
36742 
diff
changeset
 | 
104  | 
val derivation_name: thm -> string  | 
| 
 
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: 
36742 
diff
changeset
 | 
105  | 
val name_derivation: string -> thm -> thm  | 
| 
28675
 
fb68c0767004
renamed get_axiom_i to axiom, removed obsolete get_axiom;
 
wenzelm 
parents: 
28648 
diff
changeset
 | 
106  | 
val axiom: theory -> string -> thm  | 
| 
 
fb68c0767004
renamed get_axiom_i to axiom, removed obsolete get_axiom;
 
wenzelm 
parents: 
28648 
diff
changeset
 | 
107  | 
val axioms_of: theory -> (string * thm) list  | 
| 28017 | 108  | 
val get_tags: thm -> Properties.T  | 
109  | 
val map_tags: (Properties.T -> Properties.T) -> thm -> thm  | 
|
| 
23781
 
ab793a6ddf9f
Added function norm_proof for normalizing the proof term
 
berghofe 
parents: 
23657 
diff
changeset
 | 
110  | 
val norm_proof: thm -> thm  | 
| 20261 | 111  | 
val adjust_maxidx_thm: int -> thm -> thm  | 
| 38709 | 112  | 
(*meta rules*)  | 
113  | 
val assume: cterm -> thm  | 
|
114  | 
val implies_intr: cterm -> thm -> thm  | 
|
115  | 
val implies_elim: thm -> thm -> thm  | 
|
116  | 
val forall_intr: cterm -> thm -> thm  | 
|
117  | 
val forall_elim: cterm -> thm -> thm  | 
|
118  | 
val reflexive: cterm -> thm  | 
|
119  | 
val symmetric: thm -> thm  | 
|
120  | 
val transitive: thm -> thm -> thm  | 
|
121  | 
val beta_conversion: bool -> conv  | 
|
122  | 
val eta_conversion: conv  | 
|
123  | 
val eta_long_conversion: conv  | 
|
124  | 
val abstract_rule: string -> cterm -> thm -> thm  | 
|
125  | 
val combination: thm -> thm -> thm  | 
|
126  | 
val equal_intr: thm -> thm -> thm  | 
|
127  | 
val equal_elim: thm -> thm -> thm  | 
|
128  | 
val flexflex_rule: thm -> thm Seq.seq  | 
|
129  | 
val generalize: string list * string list -> int -> thm -> thm  | 
|
130  | 
val instantiate: (ctyp * ctyp) list * (cterm * cterm) list -> thm -> thm  | 
|
131  | 
val instantiate_cterm: (ctyp * ctyp) list * (cterm * cterm) list -> cterm -> cterm  | 
|
132  | 
val trivial: cterm -> thm  | 
|
| 
36330
 
0584e203960e
renamed Drule.unconstrainTs to Thm.unconstrain_allTs to accomdate the version by krauss/schropp;
 
wenzelm 
parents: 
35986 
diff
changeset
 | 
133  | 
val of_class: ctyp * class -> thm  | 
| 
36614
 
b6c031ad3690
minor tuning of Thm.strip_shyps -- no longer pervasive;
 
wenzelm 
parents: 
36613 
diff
changeset
 | 
134  | 
val strip_shyps: thm -> thm  | 
| 
36768
 
46be86127972
just one version of Thm.unconstrainT, which affects all variables;
 
wenzelm 
parents: 
36744 
diff
changeset
 | 
135  | 
val unconstrainT: thm -> thm  | 
| 38709 | 136  | 
val varifyT_global': (string * sort) list -> thm -> ((string * sort) * indexname) list * thm  | 
137  | 
val varifyT_global: thm -> thm  | 
|
| 
36615
 
88756a5a92fc
renamed Thm.freezeT to Thm.legacy_freezeT -- it is based on Type.legacy_freeze;
 
wenzelm 
parents: 
36614 
diff
changeset
 | 
138  | 
val legacy_freezeT: thm -> thm  | 
| 38709 | 139  | 
val dest_state: thm * int -> (term * term) list * term list * term * term  | 
140  | 
val lift_rule: cterm -> thm -> thm  | 
|
141  | 
val incr_indexes: int -> thm -> thm  | 
|
| 31945 | 142  | 
val assumption: int -> thm -> thm Seq.seq  | 
143  | 
val eq_assumption: int -> thm -> thm  | 
|
144  | 
val rotate_rule: int -> int -> thm -> thm  | 
|
145  | 
val permute_prems: int -> int -> thm -> thm  | 
|
146  | 
val rename_params_rule: string list * int -> thm -> thm  | 
|
| 38709 | 147  | 
val rename_boundvars: term -> term -> thm -> thm  | 
| 31945 | 148  | 
val compose_no_flatten: bool -> thm * int -> int -> thm -> thm Seq.seq  | 
149  | 
val bicompose: bool -> bool * thm * int -> int -> thm -> thm Seq.seq  | 
|
150  | 
val biresolution: bool -> (bool * thm) list -> int -> thm -> thm Seq.seq  | 
|
| 
42358
 
b47d41d9f4b5
Name_Space: proper configuration options long_names, short_names, unique_names instead of former unsynchronized references;
 
wenzelm 
parents: 
41700 
diff
changeset
 | 
151  | 
val extern_oracles: Proof.context -> xstring list  | 
| 
30288
 
a32700e45ab3
Thm.add_oracle interface: replaced old bstring by binding;
 
wenzelm 
parents: 
29636 
diff
changeset
 | 
152  | 
  val add_oracle: binding * ('a -> cterm) -> theory -> (string * ('a -> thm)) * theory
 | 
| 6089 | 153  | 
end;  | 
154  | 
||
| 
32590
 
95f4f08f950f
replaced opaque signature matching by plain old abstype (again, cf. ac4498f95d1c) -- this recovers pretty printing in SML/NJ and Poly/ML 5.3;
 
wenzelm 
parents: 
32198 
diff
changeset
 | 
155  | 
structure Thm: THM =  | 
| 0 | 156  | 
struct  | 
| 250 | 157  | 
|
| 
387
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
158  | 
(*** Certified terms and types ***)  | 
| 
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
159  | 
|
| 250 | 160  | 
(** certified types **)  | 
161  | 
||
| 
32590
 
95f4f08f950f
replaced opaque signature matching by plain old abstype (again, cf. ac4498f95d1c) -- this recovers pretty printing in SML/NJ and Poly/ML 5.3;
 
wenzelm 
parents: 
32198 
diff
changeset
 | 
162  | 
abstype ctyp = Ctyp of  | 
| 20512 | 163  | 
 {thy_ref: theory_ref,
 | 
164  | 
T: typ,  | 
|
165  | 
maxidx: int,  | 
|
| 39687 | 166  | 
sorts: sort Ord_List.T}  | 
| 
32590
 
95f4f08f950f
replaced opaque signature matching by plain old abstype (again, cf. ac4498f95d1c) -- this recovers pretty printing in SML/NJ and Poly/ML 5.3;
 
wenzelm 
parents: 
32198 
diff
changeset
 | 
167  | 
with  | 
| 250 | 168  | 
|
| 
26631
 
d6b6c74a8bcf
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
26424 
diff
changeset
 | 
169  | 
fun rep_ctyp (Ctyp args) = args;  | 
| 16656 | 170  | 
fun theory_of_ctyp (Ctyp {thy_ref, ...}) = Theory.deref thy_ref;
 | 
| 250 | 171  | 
fun typ_of (Ctyp {T, ...}) = T;
 | 
172  | 
||
| 16656 | 173  | 
fun ctyp_of thy raw_T =  | 
| 
24143
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
174  | 
let  | 
| 
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
175  | 
val T = Sign.certify_typ thy raw_T;  | 
| 
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
176  | 
val maxidx = Term.maxidx_of_typ T;  | 
| 
26640
 
92e6d3ec91bd
simplified handling of sorts, removed unnecessary universal witness;
 
wenzelm 
parents: 
26631 
diff
changeset
 | 
177  | 
val sorts = Sorts.insert_typ T [];  | 
| 
24143
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
178  | 
  in Ctyp {thy_ref = Theory.check_thy thy, T = T, maxidx = maxidx, sorts = sorts} end;
 | 
| 250 | 179  | 
|
| 32784 | 180  | 
fun dest_ctyp (Ctyp {thy_ref, T = Type (_, Ts), maxidx, sorts}) =
 | 
| 20512 | 181  | 
      map (fn T => Ctyp {thy_ref = thy_ref, T = T, maxidx = maxidx, sorts = sorts}) Ts
 | 
| 16679 | 182  | 
  | dest_ctyp cT = raise TYPE ("dest_ctyp", [typ_of cT], []);
 | 
| 15087 | 183  | 
|
| 
229
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
225 
diff
changeset
 | 
184  | 
|
| 
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
225 
diff
changeset
 | 
185  | 
|
| 250 | 186  | 
(** certified terms **)  | 
| 
229
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
225 
diff
changeset
 | 
187  | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
188  | 
(*certified terms with checked typ, maxidx, and sorts*)  | 
| 
32590
 
95f4f08f950f
replaced opaque signature matching by plain old abstype (again, cf. ac4498f95d1c) -- this recovers pretty printing in SML/NJ and Poly/ML 5.3;
 
wenzelm 
parents: 
32198 
diff
changeset
 | 
189  | 
abstype cterm = Cterm of  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
190  | 
 {thy_ref: theory_ref,
 | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
191  | 
t: term,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
192  | 
T: typ,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
193  | 
maxidx: int,  | 
| 39687 | 194  | 
sorts: sort Ord_List.T}  | 
| 
32590
 
95f4f08f950f
replaced opaque signature matching by plain old abstype (again, cf. ac4498f95d1c) -- this recovers pretty printing in SML/NJ and Poly/ML 5.3;
 
wenzelm 
parents: 
32198 
diff
changeset
 | 
195  | 
with  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
196  | 
|
| 22584 | 197  | 
exception CTERM of string * cterm list;  | 
| 16679 | 198  | 
|
| 
26631
 
d6b6c74a8bcf
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
26424 
diff
changeset
 | 
199  | 
fun rep_cterm (Cterm args) = args;  | 
| 
229
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
225 
diff
changeset
 | 
200  | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
201  | 
fun crep_cterm (Cterm {thy_ref, t, T, maxidx, sorts}) =
 | 
| 
26631
 
d6b6c74a8bcf
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
26424 
diff
changeset
 | 
202  | 
  {thy_ref = thy_ref, t = t, maxidx = maxidx, sorts = sorts,
 | 
| 
 
d6b6c74a8bcf
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
26424 
diff
changeset
 | 
203  | 
    T = Ctyp {thy_ref = thy_ref, T = T, maxidx = maxidx, sorts = sorts}};
 | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
204  | 
|
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
205  | 
fun theory_of_cterm (Cterm {thy_ref, ...}) = Theory.deref thy_ref;
 | 
| 250 | 206  | 
fun term_of (Cterm {t, ...}) = t;
 | 
| 
229
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
225 
diff
changeset
 | 
207  | 
|
| 20512 | 208  | 
fun ctyp_of_term (Cterm {thy_ref, T, maxidx, sorts, ...}) =
 | 
209  | 
  Ctyp {thy_ref = thy_ref, T = T, maxidx = maxidx, sorts = sorts};
 | 
|
| 2671 | 210  | 
|
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
211  | 
fun cterm_of thy tm =  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
212  | 
let  | 
| 18969 | 213  | 
val (t, T, maxidx) = Sign.certify_term thy tm;  | 
| 
26640
 
92e6d3ec91bd
simplified handling of sorts, removed unnecessary universal witness;
 
wenzelm 
parents: 
26631 
diff
changeset
 | 
214  | 
val sorts = Sorts.insert_term t [];  | 
| 
24143
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
215  | 
  in Cterm {thy_ref = Theory.check_thy thy, t = t, T = T, maxidx = maxidx, sorts = sorts} end;
 | 
| 
229
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
225 
diff
changeset
 | 
216  | 
|
| 32784 | 217  | 
fun merge_thys0 (Cterm {thy_ref = r1, ...}) (Cterm {thy_ref = r2, ...}) =
 | 
| 23601 | 218  | 
Theory.merge_refs (r1, r2);  | 
| 16656 | 219  | 
|
| 
20580
 
6fb75df09253
added dest_arg, i.e. a tuned version of #2 o dest_comb;
 
wenzelm 
parents: 
20548 
diff
changeset
 | 
220  | 
|
| 22909 | 221  | 
(* destructors *)  | 
222  | 
||
| 32784 | 223  | 
fun dest_comb (Cterm {t = c $ a, T, thy_ref, maxidx, sorts}) =
 | 
| 22909 | 224  | 
let val A = Term.argument_type_of c 0 in  | 
225  | 
        (Cterm {t = c, T = A --> T, thy_ref = thy_ref, maxidx = maxidx, sorts = sorts},
 | 
|
226  | 
         Cterm {t = a, T = A, thy_ref = thy_ref, maxidx = maxidx, sorts = sorts})
 | 
|
| 
1493
 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 
clasohm 
parents: 
1460 
diff
changeset
 | 
227  | 
end  | 
| 22584 | 228  | 
  | dest_comb ct = raise CTERM ("dest_comb", [ct]);
 | 
| 
1493
 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 
clasohm 
parents: 
1460 
diff
changeset
 | 
229  | 
|
| 32784 | 230  | 
fun dest_fun (Cterm {t = c $ _, T, thy_ref, maxidx, sorts}) =
 | 
| 22909 | 231  | 
let val A = Term.argument_type_of c 0  | 
232  | 
      in Cterm {t = c, T = A --> T, thy_ref = thy_ref, maxidx = maxidx, sorts = sorts} end
 | 
|
233  | 
  | dest_fun ct = raise CTERM ("dest_fun", [ct]);
 | 
|
234  | 
||
| 32784 | 235  | 
fun dest_arg (Cterm {t = c $ a, T = _, thy_ref, maxidx, sorts}) =
 | 
| 22909 | 236  | 
let val A = Term.argument_type_of c 0  | 
237  | 
      in Cterm {t = a, T = A, thy_ref = thy_ref, maxidx = maxidx, sorts = sorts} end
 | 
|
| 22584 | 238  | 
  | dest_arg ct = raise CTERM ("dest_arg", [ct]);
 | 
| 
20580
 
6fb75df09253
added dest_arg, i.e. a tuned version of #2 o dest_comb;
 
wenzelm 
parents: 
20548 
diff
changeset
 | 
239  | 
|
| 22909 | 240  | 
|
| 32784 | 241  | 
fun dest_fun2 (Cterm {t = c $ _ $ _, T, thy_ref, maxidx, sorts}) =
 | 
| 22909 | 242  | 
let  | 
243  | 
val A = Term.argument_type_of c 0;  | 
|
244  | 
val B = Term.argument_type_of c 1;  | 
|
245  | 
      in Cterm {t = c, T = A --> B --> T, thy_ref = thy_ref, maxidx = maxidx, sorts = sorts} end
 | 
|
246  | 
  | dest_fun2 ct = raise CTERM ("dest_fun2", [ct]);
 | 
|
247  | 
||
248  | 
fun dest_arg1 (Cterm {t = c $ a $ _, T = _, thy_ref, maxidx, sorts}) =
 | 
|
249  | 
let val A = Term.argument_type_of c 0  | 
|
250  | 
      in Cterm {t = a, T = A, thy_ref = thy_ref, maxidx = maxidx, sorts = sorts} end
 | 
|
251  | 
  | dest_arg1 ct = raise CTERM ("dest_arg1", [ct]);
 | 
|
| 20673 | 252  | 
|
| 32784 | 253  | 
fun dest_abs a (Cterm {t = Abs (x, T, t), T = Type ("fun", [_, U]), thy_ref, maxidx, sorts}) =
 | 
| 18944 | 254  | 
let val (y', t') = Term.dest_abs (the_default x a, T, t) in  | 
| 16679 | 255  | 
        (Cterm {t = Free (y', T), T = T, thy_ref = thy_ref, maxidx = maxidx, sorts = sorts},
 | 
256  | 
          Cterm {t = t', T = U, thy_ref = thy_ref, maxidx = maxidx, sorts = sorts})
 | 
|
| 
1493
 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 
clasohm 
parents: 
1460 
diff
changeset
 | 
257  | 
end  | 
| 22584 | 258  | 
  | dest_abs _ ct = raise CTERM ("dest_abs", [ct]);
 | 
| 
1493
 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 
clasohm 
parents: 
1460 
diff
changeset
 | 
259  | 
|
| 22909 | 260  | 
|
261  | 
(* constructors *)  | 
|
262  | 
||
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
263  | 
fun capply  | 
| 16656 | 264  | 
  (cf as Cterm {t = f, T = Type ("fun", [dty, rty]), maxidx = maxidx1, sorts = sorts1, ...})
 | 
265  | 
  (cx as Cterm {t = x, T, maxidx = maxidx2, sorts = sorts2, ...}) =
 | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
266  | 
if T = dty then  | 
| 16656 | 267  | 
      Cterm {thy_ref = merge_thys0 cf cx,
 | 
268  | 
t = f $ x,  | 
|
269  | 
T = rty,  | 
|
270  | 
maxidx = Int.max (maxidx1, maxidx2),  | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
271  | 
sorts = Sorts.union sorts1 sorts2}  | 
| 22584 | 272  | 
      else raise CTERM ("capply: types don't agree", [cf, cx])
 | 
273  | 
  | capply cf cx = raise CTERM ("capply: first arg is not a function", [cf, cx]);
 | 
|
| 250 | 274  | 
|
| 32198 | 275  | 
fun cabs_name  | 
276  | 
  (x, ct1 as Cterm {t = t1, T = T1, maxidx = maxidx1, sorts = sorts1, ...})
 | 
|
| 16656 | 277  | 
  (ct2 as Cterm {t = t2, T = T2, maxidx = maxidx2, sorts = sorts2, ...}) =
 | 
| 32198 | 278  | 
let val t = Term.lambda_name (x, t1) t2 in  | 
| 16656 | 279  | 
      Cterm {thy_ref = merge_thys0 ct1 ct2,
 | 
280  | 
t = t, T = T1 --> T2,  | 
|
281  | 
maxidx = Int.max (maxidx1, maxidx2),  | 
|
282  | 
sorts = Sorts.union sorts1 sorts2}  | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
283  | 
end;  | 
| 
229
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
225 
diff
changeset
 | 
284  | 
|
| 32198 | 285  | 
fun cabs t u = cabs_name ("", t) u;
 | 
286  | 
||
| 
20580
 
6fb75df09253
added dest_arg, i.e. a tuned version of #2 o dest_comb;
 
wenzelm 
parents: 
20548 
diff
changeset
 | 
287  | 
|
| 22909 | 288  | 
(* indexes *)  | 
289  | 
||
| 
20580
 
6fb75df09253
added dest_arg, i.e. a tuned version of #2 o dest_comb;
 
wenzelm 
parents: 
20548 
diff
changeset
 | 
290  | 
fun adjust_maxidx_cterm i (ct as Cterm {thy_ref, t, T, maxidx, sorts}) =
 | 
| 
 
6fb75df09253
added dest_arg, i.e. a tuned version of #2 o dest_comb;
 
wenzelm 
parents: 
20548 
diff
changeset
 | 
291  | 
if maxidx = i then ct  | 
| 
 
6fb75df09253
added dest_arg, i.e. a tuned version of #2 o dest_comb;
 
wenzelm 
parents: 
20548 
diff
changeset
 | 
292  | 
else if maxidx < i then  | 
| 
 
6fb75df09253
added dest_arg, i.e. a tuned version of #2 o dest_comb;
 
wenzelm 
parents: 
20548 
diff
changeset
 | 
293  | 
    Cterm {maxidx = i, thy_ref = thy_ref, t = t, T = T, sorts = sorts}
 | 
| 
 
6fb75df09253
added dest_arg, i.e. a tuned version of #2 o dest_comb;
 
wenzelm 
parents: 
20548 
diff
changeset
 | 
294  | 
else  | 
| 
 
6fb75df09253
added dest_arg, i.e. a tuned version of #2 o dest_comb;
 
wenzelm 
parents: 
20548 
diff
changeset
 | 
295  | 
    Cterm {maxidx = Int.max (maxidx_of_term t, i), thy_ref = thy_ref, t = t, T = T, sorts = sorts};
 | 
| 
 
6fb75df09253
added dest_arg, i.e. a tuned version of #2 o dest_comb;
 
wenzelm 
parents: 
20548 
diff
changeset
 | 
296  | 
|
| 22909 | 297  | 
fun incr_indexes_cterm i (ct as Cterm {thy_ref, t, T, maxidx, sorts}) =
 | 
298  | 
  if i < 0 then raise CTERM ("negative increment", [ct])
 | 
|
299  | 
else if i = 0 then ct  | 
|
300  | 
  else Cterm {thy_ref = thy_ref, t = Logic.incr_indexes ([], i) t,
 | 
|
301  | 
T = Logic.incr_tvar i T, maxidx = maxidx + i, sorts = sorts};  | 
|
302  | 
||
303  | 
||
304  | 
(* matching *)  | 
|
305  | 
||
306  | 
local  | 
|
307  | 
||
308  | 
fun gen_match match  | 
|
| 20512 | 309  | 
    (ct1 as Cterm {t = t1, sorts = sorts1, ...},
 | 
| 20815 | 310  | 
     ct2 as Cterm {t = t2, sorts = sorts2, maxidx = maxidx2, ...}) =
 | 
| 
10416
 
5b33e732e459
- Moved rewriting functions to meta_simplifier.ML
 
berghofe 
parents: 
10348 
diff
changeset
 | 
311  | 
let  | 
| 
24143
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
312  | 
val thy = Theory.deref (merge_thys0 ct1 ct2);  | 
| 
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
313  | 
val (Tinsts, tinsts) = match thy (t1, t2) (Vartab.empty, Vartab.empty);  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
314  | 
val sorts = Sorts.union sorts1 sorts2;  | 
| 20512 | 315  | 
fun mk_cTinst ((a, i), (S, T)) =  | 
| 
24143
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
316  | 
      (Ctyp {T = TVar ((a, i), S), thy_ref = Theory.check_thy thy, maxidx = i, sorts = sorts},
 | 
| 
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
317  | 
       Ctyp {T = T, thy_ref = Theory.check_thy thy, maxidx = maxidx2, sorts = sorts});
 | 
| 20512 | 318  | 
fun mk_ctinst ((x, i), (T, t)) =  | 
| 32035 | 319  | 
let val T = Envir.subst_type Tinsts T in  | 
| 
24143
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
320  | 
        (Cterm {t = Var ((x, i), T), T = T, thy_ref = Theory.check_thy thy,
 | 
| 
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
321  | 
maxidx = i, sorts = sorts},  | 
| 
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
322  | 
         Cterm {t = t, T = T, thy_ref = Theory.check_thy thy, maxidx = maxidx2, sorts = sorts})
 | 
| 
10416
 
5b33e732e459
- Moved rewriting functions to meta_simplifier.ML
 
berghofe 
parents: 
10348 
diff
changeset
 | 
323  | 
end;  | 
| 16656 | 324  | 
in (Vartab.fold (cons o mk_cTinst) Tinsts [], Vartab.fold (cons o mk_ctinst) tinsts []) end;  | 
| 
10416
 
5b33e732e459
- Moved rewriting functions to meta_simplifier.ML
 
berghofe 
parents: 
10348 
diff
changeset
 | 
325  | 
|
| 22909 | 326  | 
in  | 
| 
10416
 
5b33e732e459
- Moved rewriting functions to meta_simplifier.ML
 
berghofe 
parents: 
10348 
diff
changeset
 | 
327  | 
|
| 22909 | 328  | 
val match = gen_match Pattern.match;  | 
329  | 
val first_order_match = gen_match Pattern.first_order_match;  | 
|
330  | 
||
331  | 
end;  | 
|
| 
10416
 
5b33e732e459
- Moved rewriting functions to meta_simplifier.ML
 
berghofe 
parents: 
10348 
diff
changeset
 | 
332  | 
|
| 2509 | 333  | 
|
334  | 
||
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
335  | 
(*** Derivations and Theorems ***)  | 
| 
229
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
225 
diff
changeset
 | 
336  | 
|
| 
32590
 
95f4f08f950f
replaced opaque signature matching by plain old abstype (again, cf. ac4498f95d1c) -- this recovers pretty printing in SML/NJ and Poly/ML 5.3;
 
wenzelm 
parents: 
32198 
diff
changeset
 | 
337  | 
abstype thm = Thm of  | 
| 40124 | 338  | 
deriv * (*derivation*)  | 
339  | 
 {thy_ref: theory_ref,          (*dynamic reference to theory*)
 | 
|
340  | 
tags: Properties.T, (*additional annotations/comments*)  | 
|
341  | 
maxidx: int, (*maximum index of any Var or TVar*)  | 
|
342  | 
shyps: sort Ord_List.T, (*sort hypotheses*)  | 
|
343  | 
hyps: term Ord_List.T, (*hypotheses*)  | 
|
344  | 
tpairs: (term * term) list, (*flex-flex pairs*)  | 
|
345  | 
prop: term} (*conclusion*)  | 
|
| 28624 | 346  | 
and deriv = Deriv of  | 
| 39687 | 347  | 
 {promises: (serial * thm future) Ord_List.T,
 | 
| 37309 | 348  | 
body: Proofterm.proof_body}  | 
| 
32590
 
95f4f08f950f
replaced opaque signature matching by plain old abstype (again, cf. ac4498f95d1c) -- this recovers pretty printing in SML/NJ and Poly/ML 5.3;
 
wenzelm 
parents: 
32198 
diff
changeset
 | 
349  | 
with  | 
| 0 | 350  | 
|
| 23601 | 351  | 
type conv = cterm -> thm;  | 
352  | 
||
| 16725 | 353  | 
(*errors involving theorems*)  | 
354  | 
exception THM of string * int * thm list;  | 
|
| 
13658
 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 
berghofe 
parents: 
13642 
diff
changeset
 | 
355  | 
|
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
356  | 
fun rep_thm (Thm (_, args)) = args;  | 
| 0 | 357  | 
|
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
358  | 
fun crep_thm (Thm (_, {thy_ref, tags, maxidx, shyps, hyps, tpairs, prop})) =
 | 
| 
26631
 
d6b6c74a8bcf
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
26424 
diff
changeset
 | 
359  | 
  let fun cterm max t = Cterm {thy_ref = thy_ref, t = t, T = propT, maxidx = max, sorts = shyps} in
 | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
360  | 
   {thy_ref = thy_ref, tags = tags, maxidx = maxidx, shyps = shyps,
 | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
361  | 
hyps = map (cterm ~1) hyps,  | 
| 
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
362  | 
tpairs = map (pairself (cterm maxidx)) tpairs,  | 
| 
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
363  | 
prop = cterm maxidx prop}  | 
| 1517 | 364  | 
end;  | 
365  | 
||
| 
31947
 
7daee3bed3af
clarified strip_shyps: proper type witnesses for present sorts;
 
wenzelm 
parents: 
31945 
diff
changeset
 | 
366  | 
fun fold_terms f (Thm (_, {tpairs, prop, hyps, ...})) =
 | 
| 
 
7daee3bed3af
clarified strip_shyps: proper type witnesses for present sorts;
 
wenzelm 
parents: 
31945 
diff
changeset
 | 
367  | 
fold (fn (t, u) => f t #> f u) tpairs #> f prop #> fold f hyps;  | 
| 
 
7daee3bed3af
clarified strip_shyps: proper type witnesses for present sorts;
 
wenzelm 
parents: 
31945 
diff
changeset
 | 
368  | 
|
| 16725 | 369  | 
fun terms_of_tpairs tpairs = fold_rev (fn (t, u) => cons t o cons u) tpairs [];  | 
370  | 
||
371  | 
fun eq_tpairs ((t, u), (t', u')) = t aconv t' andalso u aconv u';  | 
|
| 18944 | 372  | 
fun union_tpairs ts us = Library.merge eq_tpairs (ts, us);  | 
| 
16884
 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 
wenzelm 
parents: 
16847 
diff
changeset
 | 
373  | 
val maxidx_tpairs = fold (fn (t, u) => Term.maxidx_term t #> Term.maxidx_term u);  | 
| 16725 | 374  | 
|
375  | 
fun attach_tpairs tpairs prop =  | 
|
376  | 
Logic.list_implies (map Logic.mk_equals tpairs, prop);  | 
|
377  | 
||
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
378  | 
fun full_prop_of (Thm (_, {tpairs, prop, ...})) = attach_tpairs tpairs prop;
 | 
| 16945 | 379  | 
|
| 39687 | 380  | 
val union_hyps = Ord_List.union Term_Ord.fast_term_ord;  | 
381  | 
val insert_hyps = Ord_List.insert Term_Ord.fast_term_ord;  | 
|
382  | 
val remove_hyps = Ord_List.remove Term_Ord.fast_term_ord;  | 
|
| 22365 | 383  | 
|
| 16945 | 384  | 
|
| 
24143
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
385  | 
(* merge theories of cterms/thms -- trivial absorption only *)  | 
| 16945 | 386  | 
|
| 32784 | 387  | 
fun merge_thys1 (Cterm {thy_ref = r1, ...}) (Thm (_, {thy_ref = r2, ...})) =
 | 
| 23601 | 388  | 
Theory.merge_refs (r1, r2);  | 
| 16945 | 389  | 
|
| 32784 | 390  | 
fun merge_thys2 (Thm (_, {thy_ref = r1, ...})) (Thm (_, {thy_ref = r2, ...})) =
 | 
| 23601 | 391  | 
Theory.merge_refs (r1, r2);  | 
| 16945 | 392  | 
|
| 0 | 393  | 
|
| 22365 | 394  | 
(* basic components *)  | 
| 16135 | 395  | 
|
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
396  | 
val theory_of_thm = Theory.deref o #thy_ref o rep_thm;  | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
397  | 
val maxidx_of = #maxidx o rep_thm;  | 
| 19910 | 398  | 
fun maxidx_thm th i = Int.max (maxidx_of th, i);  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
399  | 
val hyps_of = #hyps o rep_thm;  | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
400  | 
val prop_of = #prop o rep_thm;  | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
401  | 
val tpairs_of = #tpairs o rep_thm;  | 
| 0 | 402  | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
403  | 
val concl_of = Logic.strip_imp_concl o prop_of;  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
404  | 
val prems_of = Logic.strip_imp_prems o prop_of;  | 
| 21576 | 405  | 
val nprems_of = Logic.count_prems o prop_of;  | 
| 19305 | 406  | 
fun no_prems th = nprems_of th = 0;  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
407  | 
|
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
408  | 
fun major_prem_of th =  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
409  | 
(case prems_of th of  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
410  | 
prem :: _ => Logic.strip_assums_concl prem  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
411  | 
  | [] => raise THM ("major_prem_of: rule with no premises", 0, [th]));
 | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
412  | 
|
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
413  | 
(*the statement of any thm is a cterm*)  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
414  | 
fun cprop_of (Thm (_, {thy_ref, maxidx, shyps, prop, ...})) =
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
415  | 
  Cterm {thy_ref = thy_ref, maxidx = maxidx, T = propT, t = prop, sorts = shyps};
 | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
416  | 
|
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
417  | 
fun cprem_of (th as Thm (_, {thy_ref, maxidx, shyps, prop, ...})) i =
 | 
| 18035 | 418  | 
  Cterm {thy_ref = thy_ref, maxidx = maxidx, T = propT, sorts = shyps,
 | 
| 18145 | 419  | 
    t = Logic.nth_prem (i, prop) handle TERM _ => raise THM ("cprem_of", i, [th])};
 | 
| 18035 | 420  | 
|
| 16656 | 421  | 
(*explicit transfer to a super theory*)  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
422  | 
fun transfer thy' thm =  | 
| 3895 | 423  | 
let  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
424  | 
    val Thm (der, {thy_ref, tags, maxidx, shyps, hyps, tpairs, prop}) = thm;
 | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
425  | 
val thy = Theory.deref thy_ref;  | 
| 26665 | 426  | 
    val _ = Theory.subthy (thy, thy') orelse raise THM ("transfer: not a super theory", 0, [thm]);
 | 
427  | 
val is_eq = Theory.eq_thy (thy, thy');  | 
|
| 
24143
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
428  | 
val _ = Theory.check_thy thy;  | 
| 3895 | 429  | 
in  | 
| 
24143
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
430  | 
if is_eq then thm  | 
| 16945 | 431  | 
else  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
432  | 
Thm (der,  | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
433  | 
       {thy_ref = Theory.check_thy thy',
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
434  | 
tags = tags,  | 
| 16945 | 435  | 
maxidx = maxidx,  | 
436  | 
shyps = shyps,  | 
|
437  | 
hyps = hyps,  | 
|
438  | 
tpairs = tpairs,  | 
|
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
439  | 
prop = prop})  | 
| 3895 | 440  | 
end;  | 
| 
387
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
441  | 
|
| 16945 | 442  | 
(*explicit weakening: maps |- B to A |- B*)  | 
443  | 
fun weaken raw_ct th =  | 
|
444  | 
let  | 
|
| 20261 | 445  | 
    val ct as Cterm {t = A, T, sorts, maxidx = maxidxA, ...} = adjust_maxidx_cterm ~1 raw_ct;
 | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
446  | 
    val Thm (der, {tags, maxidx, shyps, hyps, tpairs, prop, ...}) = th;
 | 
| 16945 | 447  | 
in  | 
448  | 
if T <> propT then  | 
|
449  | 
      raise THM ("weaken: assumptions must have type prop", 0, [])
 | 
|
450  | 
else if maxidxA <> ~1 then  | 
|
451  | 
      raise THM ("weaken: assumptions may not contain schematic variables", maxidxA, [])
 | 
|
452  | 
else  | 
|
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
453  | 
Thm (der,  | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
454  | 
       {thy_ref = merge_thys1 ct th,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
455  | 
tags = tags,  | 
| 16945 | 456  | 
maxidx = maxidx,  | 
457  | 
shyps = Sorts.union sorts shyps,  | 
|
| 28354 | 458  | 
hyps = insert_hyps A hyps,  | 
| 16945 | 459  | 
tpairs = tpairs,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
460  | 
prop = prop})  | 
| 16945 | 461  | 
end;  | 
| 16656 | 462  | 
|
| 28624 | 463  | 
fun weaken_sorts raw_sorts ct =  | 
464  | 
let  | 
|
465  | 
    val Cterm {thy_ref, t, T, maxidx, sorts} = ct;
 | 
|
466  | 
val thy = Theory.deref thy_ref;  | 
|
467  | 
val more_sorts = Sorts.make (map (Sign.certify_sort thy) raw_sorts);  | 
|
468  | 
val sorts' = Sorts.union sorts more_sorts;  | 
|
469  | 
  in Cterm {thy_ref = Theory.check_thy thy, t = t, T = T, maxidx = maxidx, sorts = sorts'} end;
 | 
|
470  | 
||
| 16656 | 471  | 
(*dangling sort constraints of a thm*)  | 
| 
31947
 
7daee3bed3af
clarified strip_shyps: proper type witnesses for present sorts;
 
wenzelm 
parents: 
31945 
diff
changeset
 | 
472  | 
fun extra_shyps (th as Thm (_, {shyps, ...})) =
 | 
| 
 
7daee3bed3af
clarified strip_shyps: proper type witnesses for present sorts;
 
wenzelm 
parents: 
31945 
diff
changeset
 | 
473  | 
Sorts.subtract (fold_terms Sorts.insert_term th []) shyps;  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
474  | 
|
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
475  | 
|
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
476  | 
|
| 32725 | 477  | 
(** derivations and promised proofs **)  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
478  | 
|
| 
32059
 
7991cdf10a54
support for arbitrarity nested future proofs -- replaced crude order by explicit normalization (which might loop for bad dependencies);
 
wenzelm 
parents: 
32035 
diff
changeset
 | 
479  | 
fun make_deriv promises oracles thms proof =  | 
| 
 
7991cdf10a54
support for arbitrarity nested future proofs -- replaced crude order by explicit normalization (which might loop for bad dependencies);
 
wenzelm 
parents: 
32035 
diff
changeset
 | 
480  | 
  Deriv {promises = promises, body = PBody {oracles = oracles, thms = thms, proof = proof}};
 | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
481  | 
|
| 37309 | 482  | 
val empty_deriv = make_deriv [] [] [] Proofterm.MinProof;  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
483  | 
|
| 28330 | 484  | 
|
| 28354 | 485  | 
(* inference rules *)  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
486  | 
|
| 
28378
 
60cc0359363d
promise: include check into future body, i.e. joined results are always valid;
 
wenzelm 
parents: 
28364 
diff
changeset
 | 
487  | 
fun promise_ord ((i, _), (j, _)) = int_ord (j, i);  | 
| 28330 | 488  | 
|
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
489  | 
fun deriv_rule2 f  | 
| 
32059
 
7991cdf10a54
support for arbitrarity nested future proofs -- replaced crude order by explicit normalization (which might loop for bad dependencies);
 
wenzelm 
parents: 
32035 
diff
changeset
 | 
490  | 
    (Deriv {promises = ps1, body = PBody {oracles = oras1, thms = thms1, proof = prf1}})
 | 
| 
 
7991cdf10a54
support for arbitrarity nested future proofs -- replaced crude order by explicit normalization (which might loop for bad dependencies);
 
wenzelm 
parents: 
32035 
diff
changeset
 | 
491  | 
    (Deriv {promises = ps2, body = PBody {oracles = oras2, thms = thms2, proof = prf2}}) =
 | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
492  | 
let  | 
| 39687 | 493  | 
val ps = Ord_List.union promise_ord ps1 ps2;  | 
| 44334 | 494  | 
val oras = Proofterm.unions_oracles [oras1, oras2];  | 
495  | 
val thms = Proofterm.unions_thms [thms1, thms2];  | 
|
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
496  | 
val prf =  | 
| 37309 | 497  | 
(case ! Proofterm.proofs of  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
498  | 
2 => f prf1 prf2  | 
| 
28804
 
5d3b1df16353
refined notion of derivation, consiting of promises and proof_body;
 
wenzelm 
parents: 
28675 
diff
changeset
 | 
499  | 
| 1 => MinProof  | 
| 
 
5d3b1df16353
refined notion of derivation, consiting of promises and proof_body;
 
wenzelm 
parents: 
28675 
diff
changeset
 | 
500  | 
| 0 => MinProof  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
501  | 
      | i => error ("Illegal level of detail for proof objects: " ^ string_of_int i));
 | 
| 
32059
 
7991cdf10a54
support for arbitrarity nested future proofs -- replaced crude order by explicit normalization (which might loop for bad dependencies);
 
wenzelm 
parents: 
32035 
diff
changeset
 | 
502  | 
in make_deriv ps oras thms prf end;  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
503  | 
|
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
504  | 
fun deriv_rule1 f = deriv_rule2 (K f) empty_deriv;  | 
| 
32059
 
7991cdf10a54
support for arbitrarity nested future proofs -- replaced crude order by explicit normalization (which might loop for bad dependencies);
 
wenzelm 
parents: 
32035 
diff
changeset
 | 
505  | 
fun deriv_rule0 prf = deriv_rule1 I (make_deriv [] [] [] prf);  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
506  | 
|
| 
36621
 
2fd4e2c76636
proof terms for strip_shyps, based on the version by krauss/schropp with some notable differences:
 
wenzelm 
parents: 
36619 
diff
changeset
 | 
507  | 
fun deriv_rule_unconditional f (Deriv {promises, body = PBody {oracles, thms, proof}}) =
 | 
| 
 
2fd4e2c76636
proof terms for strip_shyps, based on the version by krauss/schropp with some notable differences:
 
wenzelm 
parents: 
36619 
diff
changeset
 | 
508  | 
make_deriv promises oracles thms (f proof);  | 
| 
 
2fd4e2c76636
proof terms for strip_shyps, based on the version by krauss/schropp with some notable differences:
 
wenzelm 
parents: 
36619 
diff
changeset
 | 
509  | 
|
| 1238 | 510  | 
|
| 32725 | 511  | 
(* fulfilled proofs *)  | 
512  | 
||
| 44331 | 513  | 
fun raw_body_of (Thm (Deriv {body, ...}, _)) = body;
 | 
514  | 
fun raw_promises_of (Thm (Deriv {promises, ...}, _)) = promises;
 | 
|
515  | 
||
516  | 
fun join_promises [] = ()  | 
|
517  | 
| join_promises promises = join_promises_of (Future.joins (map snd promises))  | 
|
518  | 
and join_promises_of thms = join_promises (maps raw_promises_of thms);  | 
|
| 32725 | 519  | 
|
520  | 
fun fulfill_body (Thm (Deriv {promises, body}, {thy_ref, ...})) =
 | 
|
| 44331 | 521  | 
Proofterm.fulfill_norm_proof (Theory.deref thy_ref) (fulfill_promises promises) body  | 
522  | 
and fulfill_promises promises =  | 
|
523  | 
map fst promises ~~ map fulfill_body (Future.joins (map snd promises));  | 
|
| 32725 | 524  | 
|
| 
44333
 
cc53ce50f738
reverted to join_bodies/join_proofs based on fold_body_thms to regain performance (escpecially of HOL-Proofs) -- see also aa9c1e9ef2ce and 4e2abb045eac;
 
wenzelm 
parents: 
44331 
diff
changeset
 | 
525  | 
fun proof_bodies_of thms =  | 
| 
 
cc53ce50f738
reverted to join_bodies/join_proofs based on fold_body_thms to regain performance (escpecially of HOL-Proofs) -- see also aa9c1e9ef2ce and 4e2abb045eac;
 
wenzelm 
parents: 
44331 
diff
changeset
 | 
526  | 
let  | 
| 
 
cc53ce50f738
reverted to join_bodies/join_proofs based on fold_body_thms to regain performance (escpecially of HOL-Proofs) -- see also aa9c1e9ef2ce and 4e2abb045eac;
 
wenzelm 
parents: 
44331 
diff
changeset
 | 
527  | 
val _ = join_promises_of thms;  | 
| 
 
cc53ce50f738
reverted to join_bodies/join_proofs based on fold_body_thms to regain performance (escpecially of HOL-Proofs) -- see also aa9c1e9ef2ce and 4e2abb045eac;
 
wenzelm 
parents: 
44331 
diff
changeset
 | 
528  | 
val bodies = map fulfill_body thms;  | 
| 
 
cc53ce50f738
reverted to join_bodies/join_proofs based on fold_body_thms to regain performance (escpecially of HOL-Proofs) -- see also aa9c1e9ef2ce and 4e2abb045eac;
 
wenzelm 
parents: 
44331 
diff
changeset
 | 
529  | 
val _ = Proofterm.join_bodies bodies;  | 
| 
 
cc53ce50f738
reverted to join_bodies/join_proofs based on fold_body_thms to regain performance (escpecially of HOL-Proofs) -- see also aa9c1e9ef2ce and 4e2abb045eac;
 
wenzelm 
parents: 
44331 
diff
changeset
 | 
530  | 
in bodies end;  | 
| 
 
cc53ce50f738
reverted to join_bodies/join_proofs based on fold_body_thms to regain performance (escpecially of HOL-Proofs) -- see also aa9c1e9ef2ce and 4e2abb045eac;
 
wenzelm 
parents: 
44331 
diff
changeset
 | 
531  | 
|
| 
 
cc53ce50f738
reverted to join_bodies/join_proofs based on fold_body_thms to regain performance (escpecially of HOL-Proofs) -- see also aa9c1e9ef2ce and 4e2abb045eac;
 
wenzelm 
parents: 
44331 
diff
changeset
 | 
532  | 
val proof_body_of = singleton proof_bodies_of;  | 
| 44331 | 533  | 
val proof_of = Proofterm.proof_of o proof_body_of;  | 
| 32725 | 534  | 
|
| 
44333
 
cc53ce50f738
reverted to join_bodies/join_proofs based on fold_body_thms to regain performance (escpecially of HOL-Proofs) -- see also aa9c1e9ef2ce and 4e2abb045eac;
 
wenzelm 
parents: 
44331 
diff
changeset
 | 
535  | 
val join_proofs = ignore o proof_bodies_of;  | 
| 32725 | 536  | 
|
537  | 
||
538  | 
(* derivation status *)  | 
|
539  | 
||
540  | 
fun status_of (Thm (Deriv {promises, body}, _)) =
 | 
|
541  | 
let  | 
|
542  | 
val ps = map (Future.peek o snd) promises;  | 
|
543  | 
val bodies = body ::  | 
|
| 44331 | 544  | 
map_filter (fn SOME (Exn.Res th) => SOME (raw_body_of th) | _ => NONE) ps;  | 
| 37309 | 545  | 
    val {oracle, unfinished, failed} = Proofterm.status_of bodies;
 | 
| 32725 | 546  | 
in  | 
547  | 
   {oracle = oracle,
 | 
|
548  | 
unfinished = unfinished orelse exists is_none ps,  | 
|
549  | 
failed = failed orelse exists (fn SOME (Exn.Exn _) => true | _ => false) ps}  | 
|
550  | 
end;  | 
|
551  | 
||
552  | 
||
553  | 
(* future rule *)  | 
|
554  | 
||
| 
36613
 
f3157c288aca
simplified primitive Thm.future: more direct theory check, do not hardwire strip_shyps here;
 
wenzelm 
parents: 
36330 
diff
changeset
 | 
555  | 
fun future_result i orig_thy orig_shyps orig_prop thm =  | 
| 32725 | 556  | 
let  | 
| 
36613
 
f3157c288aca
simplified primitive Thm.future: more direct theory check, do not hardwire strip_shyps here;
 
wenzelm 
parents: 
36330 
diff
changeset
 | 
557  | 
    fun err msg = raise THM ("future_result: " ^ msg, 0, [thm]);
 | 
| 
 
f3157c288aca
simplified primitive Thm.future: more direct theory check, do not hardwire strip_shyps here;
 
wenzelm 
parents: 
36330 
diff
changeset
 | 
558  | 
    val Thm (Deriv {promises, ...}, {thy_ref, shyps, hyps, tpairs, prop, ...}) = thm;
 | 
| 
 
f3157c288aca
simplified primitive Thm.future: more direct theory check, do not hardwire strip_shyps here;
 
wenzelm 
parents: 
36330 
diff
changeset
 | 
559  | 
|
| 
 
f3157c288aca
simplified primitive Thm.future: more direct theory check, do not hardwire strip_shyps here;
 
wenzelm 
parents: 
36330 
diff
changeset
 | 
560  | 
val _ = Theory.eq_thy (Theory.deref thy_ref, orig_thy) orelse err "bad theory";  | 
| 32725 | 561  | 
val _ = Theory.check_thy orig_thy;  | 
562  | 
val _ = prop aconv orig_prop orelse err "bad prop";  | 
|
563  | 
val _ = null tpairs orelse err "bad tpairs";  | 
|
564  | 
val _ = null hyps orelse err "bad hyps";  | 
|
565  | 
val _ = Sorts.subset (shyps, orig_shyps) orelse err "bad shyps";  | 
|
566  | 
val _ = forall (fn (j, _) => i <> j) promises orelse err "bad dependencies";  | 
|
| 44331 | 567  | 
val _ = join_promises promises;  | 
| 32725 | 568  | 
in thm end;  | 
569  | 
||
570  | 
fun future future_thm ct =  | 
|
571  | 
let  | 
|
572  | 
    val Cterm {thy_ref = thy_ref, t = prop, T, maxidx, sorts} = ct;
 | 
|
573  | 
val thy = Context.reject_draft (Theory.deref thy_ref);  | 
|
574  | 
    val _ = T <> propT andalso raise CTERM ("future: prop expected", [ct]);
 | 
|
575  | 
||
576  | 
val i = serial ();  | 
|
577  | 
val future = future_thm |> Future.map (future_result i thy sorts prop);  | 
|
578  | 
in  | 
|
| 37309 | 579  | 
Thm (make_deriv [(i, future)] [] [] (Proofterm.promise_proof thy i prop),  | 
| 32725 | 580  | 
     {thy_ref = thy_ref,
 | 
581  | 
tags = [],  | 
|
582  | 
maxidx = maxidx,  | 
|
583  | 
shyps = sorts,  | 
|
584  | 
hyps = [],  | 
|
585  | 
tpairs = [],  | 
|
586  | 
prop = prop})  | 
|
587  | 
end;  | 
|
588  | 
||
589  | 
||
590  | 
(* closed derivations with official name *)  | 
|
591  | 
||
| 41699 | 592  | 
(*non-deterministic, depends on unknown promises*)  | 
| 37297 | 593  | 
fun derivation_name (Thm (Deriv {body, ...}, {shyps, hyps, prop, ...})) =
 | 
| 37309 | 594  | 
Proofterm.get_name shyps hyps prop (Proofterm.proof_of body);  | 
| 32725 | 595  | 
|
| 
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: 
36742 
diff
changeset
 | 
596  | 
fun name_derivation name (thm as Thm (der, args)) =  | 
| 32725 | 597  | 
let  | 
598  | 
    val Deriv {promises, body} = der;
 | 
|
| 
36882
 
f33760bb8ca0
conditionally unconstrain thm proofs -- loosely based on the version by krauss/schropp;
 
wenzelm 
parents: 
36770 
diff
changeset
 | 
599  | 
    val {thy_ref, shyps, hyps, prop, tpairs, ...} = args;
 | 
| 32725 | 600  | 
    val _ = null tpairs orelse raise THM ("put_name: unsolved flex-flex constraints", 0, [thm]);
 | 
601  | 
||
| 
41700
 
f33d5a00c25d
thm_proof: visible fulfill_body only, without joining nested thms -- retain proof irrelevance, which is important for parallel performance;
 
wenzelm 
parents: 
41699 
diff
changeset
 | 
602  | 
val ps = map (apsnd (Future.map fulfill_body)) promises;  | 
| 32725 | 603  | 
val thy = Theory.deref thy_ref;  | 
| 37309 | 604  | 
val (pthm, proof) = Proofterm.thm_proof thy name shyps hyps prop ps body;  | 
| 32725 | 605  | 
val der' = make_deriv [] [] [pthm] proof;  | 
606  | 
val _ = Theory.check_thy thy;  | 
|
607  | 
in Thm (der', args) end;  | 
|
608  | 
||
609  | 
||
| 1238 | 610  | 
|
| 1529 | 611  | 
(** Axioms **)  | 
| 
387
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
612  | 
|
| 
28675
 
fb68c0767004
renamed get_axiom_i to axiom, removed obsolete get_axiom;
 
wenzelm 
parents: 
28648 
diff
changeset
 | 
613  | 
fun axiom theory name =  | 
| 
387
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
614  | 
let  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
615  | 
fun get_ax thy =  | 
| 22685 | 616  | 
Symtab.lookup (Theory.axiom_table thy) name  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
617  | 
|> Option.map (fn prop =>  | 
| 
24143
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
618  | 
let  | 
| 37309 | 619  | 
val der = deriv_rule0 (Proofterm.axm_proof name prop);  | 
| 
24143
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
620  | 
val maxidx = maxidx_of_term prop;  | 
| 
26640
 
92e6d3ec91bd
simplified handling of sorts, removed unnecessary universal witness;
 
wenzelm 
parents: 
26631 
diff
changeset
 | 
621  | 
val shyps = Sorts.insert_term prop [];  | 
| 
24143
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
622  | 
in  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
623  | 
             Thm (der, {thy_ref = Theory.check_thy thy, tags = [],
 | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
624  | 
maxidx = maxidx, shyps = shyps, hyps = [], tpairs = [], prop = prop})  | 
| 
24143
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
625  | 
end);  | 
| 
387
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
626  | 
in  | 
| 42425 | 627  | 
(case get_first get_ax (Theory.nodes_of theory) of  | 
| 15531 | 628  | 
SOME thm => thm  | 
629  | 
    | NONE => raise THEORY ("No axiom " ^ quote name, [theory]))
 | 
|
| 
387
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
630  | 
end;  | 
| 
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
631  | 
|
| 
776
 
df8f91c0e57c
improved axioms_of: returns thms as the manual says;
 
wenzelm 
parents: 
721 
diff
changeset
 | 
632  | 
(*return additional axioms of this theory node*)  | 
| 
 
df8f91c0e57c
improved axioms_of: returns thms as the manual says;
 
wenzelm 
parents: 
721 
diff
changeset
 | 
633  | 
fun axioms_of thy =  | 
| 
28675
 
fb68c0767004
renamed get_axiom_i to axiom, removed obsolete get_axiom;
 
wenzelm 
parents: 
28648 
diff
changeset
 | 
634  | 
map (fn s => (s, axiom thy s)) (Symtab.keys (Theory.axiom_table thy));  | 
| 
776
 
df8f91c0e57c
improved axioms_of: returns thms as the manual says;
 
wenzelm 
parents: 
721 
diff
changeset
 | 
635  | 
|
| 6089 | 636  | 
|
| 
28804
 
5d3b1df16353
refined notion of derivation, consiting of promises and proof_body;
 
wenzelm 
parents: 
28675 
diff
changeset
 | 
637  | 
(* tags *)  | 
| 6089 | 638  | 
|
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
639  | 
val get_tags = #tags o rep_thm;  | 
| 6089 | 640  | 
|
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
641  | 
fun map_tags f (Thm (der, {thy_ref, tags, maxidx, shyps, hyps, tpairs, prop})) =
 | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
642  | 
  Thm (der, {thy_ref = thy_ref, tags = f tags, maxidx = maxidx,
 | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
643  | 
shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop});  | 
| 0 | 644  | 
|
645  | 
||
| 
43795
 
ca5896a836ba
recovered some runtime sharing from d6b6c74a8bcf, without the global memory bottleneck;
 
wenzelm 
parents: 
43761 
diff
changeset
 | 
646  | 
(* technical adjustments *)  | 
| 
 
ca5896a836ba
recovered some runtime sharing from d6b6c74a8bcf, without the global memory bottleneck;
 
wenzelm 
parents: 
43761 
diff
changeset
 | 
647  | 
|
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
648  | 
fun norm_proof (Thm (der, args as {thy_ref, ...})) =
 | 
| 
24143
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
649  | 
let  | 
| 
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
650  | 
val thy = Theory.deref thy_ref;  | 
| 37309 | 651  | 
val der' = deriv_rule1 (Proofterm.rew_proof thy) der;  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
652  | 
val _ = Theory.check_thy thy;  | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
653  | 
in Thm (der', args) end;  | 
| 
23781
 
ab793a6ddf9f
Added function norm_proof for normalizing the proof term
 
berghofe 
parents: 
23657 
diff
changeset
 | 
654  | 
|
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
655  | 
fun adjust_maxidx_thm i (th as Thm (der, {thy_ref, tags, maxidx, shyps, hyps, tpairs, prop})) =
 | 
| 20261 | 656  | 
if maxidx = i then th  | 
657  | 
else if maxidx < i then  | 
|
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
658  | 
    Thm (der, {maxidx = i, thy_ref = thy_ref, tags = tags, shyps = shyps,
 | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
659  | 
hyps = hyps, tpairs = tpairs, prop = prop})  | 
| 20261 | 660  | 
else  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
661  | 
    Thm (der, {maxidx = Int.max (maxidx_tpairs tpairs (maxidx_of_term prop), i), thy_ref = thy_ref,
 | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
662  | 
tags = tags, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop});  | 
| 564 | 663  | 
|
| 
387
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
664  | 
|
| 2509 | 665  | 
|
| 1529 | 666  | 
(*** Meta rules ***)  | 
| 0 | 667  | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
668  | 
(** primitive rules **)  | 
| 0 | 669  | 
|
| 16656 | 670  | 
(*The assumption rule A |- A*)  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
671  | 
fun assume raw_ct =  | 
| 20261 | 672  | 
  let val Cterm {thy_ref, t = prop, T, maxidx, sorts} = adjust_maxidx_cterm ~1 raw_ct in
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
673  | 
if T <> propT then  | 
| 19230 | 674  | 
      raise THM ("assume: prop", 0, [])
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
675  | 
else if maxidx <> ~1 then  | 
| 19230 | 676  | 
      raise THM ("assume: variables", maxidx, [])
 | 
| 37309 | 677  | 
else Thm (deriv_rule0 (Proofterm.Hyp prop),  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
678  | 
     {thy_ref = thy_ref,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
679  | 
tags = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
680  | 
maxidx = ~1,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
681  | 
shyps = sorts,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
682  | 
hyps = [prop],  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
683  | 
tpairs = [],  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
684  | 
prop = prop})  | 
| 0 | 685  | 
end;  | 
686  | 
||
| 1220 | 687  | 
(*Implication introduction  | 
| 3529 | 688  | 
[A]  | 
689  | 
:  | 
|
690  | 
B  | 
|
| 1220 | 691  | 
-------  | 
692  | 
A ==> B  | 
|
693  | 
*)  | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
694  | 
fun implies_intr  | 
| 16679 | 695  | 
    (ct as Cterm {t = A, T, maxidx = maxidxA, sorts, ...})
 | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
696  | 
    (th as Thm (der, {maxidx, hyps, shyps, tpairs, prop, ...})) =
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
697  | 
if T <> propT then  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
698  | 
    raise THM ("implies_intr: assumptions must have type prop", 0, [th])
 | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
699  | 
else  | 
| 37309 | 700  | 
Thm (deriv_rule1 (Proofterm.implies_intr_proof A) der,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
701  | 
     {thy_ref = merge_thys1 ct th,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
702  | 
tags = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
703  | 
maxidx = Int.max (maxidxA, maxidx),  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
704  | 
shyps = Sorts.union sorts shyps,  | 
| 28354 | 705  | 
hyps = remove_hyps A hyps,  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
706  | 
tpairs = tpairs,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
707  | 
prop = Logic.mk_implies (A, prop)});  | 
| 0 | 708  | 
|
| 1529 | 709  | 
|
| 1220 | 710  | 
(*Implication elimination  | 
711  | 
A ==> B A  | 
|
712  | 
------------  | 
|
713  | 
B  | 
|
714  | 
*)  | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
715  | 
fun implies_elim thAB thA =  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
716  | 
let  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
717  | 
    val Thm (derA, {maxidx = maxA, hyps = hypsA, shyps = shypsA, tpairs = tpairsA,
 | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
718  | 
prop = propA, ...}) = thA  | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
719  | 
    and Thm (der, {maxidx, hyps, shyps, tpairs, prop, ...}) = thAB;
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
720  | 
    fun err () = raise THM ("implies_elim: major premise", 0, [thAB, thA]);
 | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
721  | 
in  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
722  | 
case prop of  | 
| 20512 | 723  | 
      Const ("==>", _) $ A $ B =>
 | 
724  | 
if A aconv propA then  | 
|
| 37309 | 725  | 
Thm (deriv_rule2 (curry Proofterm.%%) der derA,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
726  | 
           {thy_ref = merge_thys2 thAB thA,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
727  | 
tags = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
728  | 
maxidx = Int.max (maxA, maxidx),  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
729  | 
shyps = Sorts.union shypsA shyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
730  | 
hyps = union_hyps hypsA hyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
731  | 
tpairs = union_tpairs tpairsA tpairs,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
732  | 
prop = B})  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
733  | 
else err ()  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
734  | 
| _ => err ()  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
735  | 
end;  | 
| 250 | 736  | 
|
| 1220 | 737  | 
(*Forall introduction. The Free or Var x must not be free in the hypotheses.  | 
| 16656 | 738  | 
[x]  | 
739  | 
:  | 
|
740  | 
A  | 
|
741  | 
------  | 
|
742  | 
!!x. A  | 
|
| 1220 | 743  | 
*)  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
744  | 
fun forall_intr  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
745  | 
    (ct as Cterm {t = x, T, sorts, ...})
 | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
746  | 
    (th as Thm (der, {maxidx, shyps, hyps, tpairs, prop, ...})) =
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
747  | 
let  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
748  | 
fun result a =  | 
| 37309 | 749  | 
Thm (deriv_rule1 (Proofterm.forall_intr_proof x a) der,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
750  | 
       {thy_ref = merge_thys1 ct th,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
751  | 
tags = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
752  | 
maxidx = maxidx,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
753  | 
shyps = Sorts.union sorts shyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
754  | 
hyps = hyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
755  | 
tpairs = tpairs,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
756  | 
prop = Term.all T $ Abs (a, T, abstract_over (x, prop))});  | 
| 21798 | 757  | 
fun check_occs a x ts =  | 
| 
16847
 
8fc160b12e73
invoke_oracle: do not keep theory value, but theory_ref;
 
wenzelm 
parents: 
16725 
diff
changeset
 | 
758  | 
if exists (fn t => Logic.occs (x, t)) ts then  | 
| 21798 | 759  | 
        raise THM ("forall_intr: variable " ^ quote a ^ " free in assumptions", 0, [th])
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
760  | 
else ();  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
761  | 
in  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
762  | 
case x of  | 
| 21798 | 763  | 
Free (a, _) => (check_occs a x hyps; check_occs a x (terms_of_tpairs tpairs); result a)  | 
764  | 
| Var ((a, _), _) => (check_occs a x (terms_of_tpairs tpairs); result a)  | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
765  | 
    | _ => raise THM ("forall_intr: not a variable", 0, [th])
 | 
| 0 | 766  | 
end;  | 
767  | 
||
| 1220 | 768  | 
(*Forall elimination  | 
| 16656 | 769  | 
!!x. A  | 
| 1220 | 770  | 
------  | 
771  | 
A[t/x]  | 
|
772  | 
*)  | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
773  | 
fun forall_elim  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
774  | 
    (ct as Cterm {t, T, maxidx = maxt, sorts, ...})
 | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
775  | 
    (th as Thm (der, {maxidx, shyps, hyps, tpairs, prop, ...})) =
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
776  | 
(case prop of  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
777  | 
    Const ("all", Type ("fun", [Type ("fun", [qary, _]), _])) $ A =>
 | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
778  | 
if T <> qary then  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
779  | 
        raise THM ("forall_elim: type mismatch", 0, [th])
 | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
780  | 
else  | 
| 37309 | 781  | 
Thm (deriv_rule1 (Proofterm.% o rpair (SOME t)) der,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
782  | 
         {thy_ref = merge_thys1 ct th,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
783  | 
tags = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
784  | 
maxidx = Int.max (maxidx, maxt),  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
785  | 
shyps = Sorts.union sorts shyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
786  | 
hyps = hyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
787  | 
tpairs = tpairs,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
788  | 
prop = Term.betapply (A, t)})  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
789  | 
  | _ => raise THM ("forall_elim: not quantified", 0, [th]));
 | 
| 0 | 790  | 
|
791  | 
||
| 1220 | 792  | 
(* Equality *)  | 
| 0 | 793  | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
794  | 
(*Reflexivity  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
795  | 
t == t  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
796  | 
*)  | 
| 32784 | 797  | 
fun reflexive (Cterm {thy_ref, t, T = _, maxidx, sorts}) =
 | 
| 37309 | 798  | 
Thm (deriv_rule0 Proofterm.reflexive,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
799  | 
   {thy_ref = thy_ref,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
800  | 
tags = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
801  | 
maxidx = maxidx,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
802  | 
shyps = sorts,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
803  | 
hyps = [],  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
804  | 
tpairs = [],  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
805  | 
prop = Logic.mk_equals (t, t)});  | 
| 0 | 806  | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
807  | 
(*Symmetry  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
808  | 
t == u  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
809  | 
------  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
810  | 
u == t  | 
| 1220 | 811  | 
*)  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
812  | 
fun symmetric (th as Thm (der, {thy_ref, maxidx, shyps, hyps, tpairs, prop, ...})) =
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
813  | 
(case prop of  | 
| 32784 | 814  | 
    (eq as Const ("==", _)) $ t $ u =>
 | 
| 37309 | 815  | 
Thm (deriv_rule1 Proofterm.symmetric der,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
816  | 
       {thy_ref = thy_ref,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
817  | 
tags = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
818  | 
maxidx = maxidx,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
819  | 
shyps = shyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
820  | 
hyps = hyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
821  | 
tpairs = tpairs,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
822  | 
prop = eq $ u $ t})  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
823  | 
    | _ => raise THM ("symmetric", 0, [th]));
 | 
| 0 | 824  | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
825  | 
(*Transitivity  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
826  | 
t1 == u u == t2  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
827  | 
------------------  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
828  | 
t1 == t2  | 
| 1220 | 829  | 
*)  | 
| 0 | 830  | 
fun transitive th1 th2 =  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
831  | 
let  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
832  | 
    val Thm (der1, {maxidx = max1, hyps = hyps1, shyps = shyps1, tpairs = tpairs1,
 | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
833  | 
prop = prop1, ...}) = th1  | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
834  | 
    and Thm (der2, {maxidx = max2, hyps = hyps2, shyps = shyps2, tpairs = tpairs2,
 | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
835  | 
prop = prop2, ...}) = th2;  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
836  | 
    fun err msg = raise THM ("transitive: " ^ msg, 0, [th1, th2]);
 | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
837  | 
in  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
838  | 
case (prop1, prop2) of  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
839  | 
      ((eq as Const ("==", Type (_, [T, _]))) $ t1 $ u, Const ("==", _) $ u' $ t2) =>
 | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
840  | 
if not (u aconv u') then err "middle term"  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
841  | 
else  | 
| 37309 | 842  | 
Thm (deriv_rule2 (Proofterm.transitive u T) der1 der2,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
843  | 
           {thy_ref = merge_thys2 th1 th2,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
844  | 
tags = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
845  | 
maxidx = Int.max (max1, max2),  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
846  | 
shyps = Sorts.union shyps1 shyps2,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
847  | 
hyps = union_hyps hyps1 hyps2,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
848  | 
tpairs = union_tpairs tpairs1 tpairs2,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
849  | 
prop = eq $ t1 $ t2})  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
850  | 
| _ => err "premises"  | 
| 0 | 851  | 
end;  | 
852  | 
||
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
853  | 
(*Beta-conversion  | 
| 16656 | 854  | 
(%x. t)(u) == t[u/x]  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
855  | 
fully beta-reduces the term if full = true  | 
| 
10416
 
5b33e732e459
- Moved rewriting functions to meta_simplifier.ML
 
berghofe 
parents: 
10348 
diff
changeset
 | 
856  | 
*)  | 
| 32784 | 857  | 
fun beta_conversion full (Cterm {thy_ref, t, T = _, maxidx, sorts}) =
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
858  | 
let val t' =  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
859  | 
if full then Envir.beta_norm t  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
860  | 
else  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
861  | 
(case t of Abs (_, _, bodt) $ u => subst_bound (u, bodt)  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
862  | 
      | _ => raise THM ("beta_conversion: not a redex", 0, []));
 | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
863  | 
in  | 
| 37309 | 864  | 
Thm (deriv_rule0 Proofterm.reflexive,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
865  | 
     {thy_ref = thy_ref,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
866  | 
tags = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
867  | 
maxidx = maxidx,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
868  | 
shyps = sorts,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
869  | 
hyps = [],  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
870  | 
tpairs = [],  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
871  | 
prop = Logic.mk_equals (t, t')})  | 
| 
10416
 
5b33e732e459
- Moved rewriting functions to meta_simplifier.ML
 
berghofe 
parents: 
10348 
diff
changeset
 | 
872  | 
end;  | 
| 
 
5b33e732e459
- Moved rewriting functions to meta_simplifier.ML
 
berghofe 
parents: 
10348 
diff
changeset
 | 
873  | 
|
| 32784 | 874  | 
fun eta_conversion (Cterm {thy_ref, t, T = _, maxidx, sorts}) =
 | 
| 37309 | 875  | 
Thm (deriv_rule0 Proofterm.reflexive,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
876  | 
   {thy_ref = thy_ref,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
877  | 
tags = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
878  | 
maxidx = maxidx,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
879  | 
shyps = sorts,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
880  | 
hyps = [],  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
881  | 
tpairs = [],  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
882  | 
prop = Logic.mk_equals (t, Envir.eta_contract t)});  | 
| 0 | 883  | 
|
| 32784 | 884  | 
fun eta_long_conversion (Cterm {thy_ref, t, T = _, maxidx, sorts}) =
 | 
| 37309 | 885  | 
Thm (deriv_rule0 Proofterm.reflexive,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
886  | 
   {thy_ref = thy_ref,
 | 
| 23493 | 887  | 
tags = [],  | 
888  | 
maxidx = maxidx,  | 
|
889  | 
shyps = sorts,  | 
|
890  | 
hyps = [],  | 
|
891  | 
tpairs = [],  | 
|
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
892  | 
prop = Logic.mk_equals (t, Pattern.eta_long [] t)});  | 
| 23493 | 893  | 
|
| 0 | 894  | 
(*The abstraction rule. The Free or Var x must not be free in the hypotheses.  | 
895  | 
The bound variable will be named "a" (since x will be something like x320)  | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
896  | 
t == u  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
897  | 
--------------  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
898  | 
%x. t == %x. u  | 
| 1220 | 899  | 
*)  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
900  | 
fun abstract_rule a  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
901  | 
    (Cterm {t = x, T, sorts, ...})
 | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
902  | 
    (th as Thm (der, {thy_ref, maxidx, hyps, shyps, tpairs, prop, ...})) =
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
903  | 
let  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
904  | 
val (t, u) = Logic.dest_equals prop  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
905  | 
      handle TERM _ => raise THM ("abstract_rule: premise not an equality", 0, [th]);
 | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
906  | 
val result =  | 
| 37309 | 907  | 
Thm (deriv_rule1 (Proofterm.abstract_rule x a) der,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
908  | 
       {thy_ref = thy_ref,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
909  | 
tags = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
910  | 
maxidx = maxidx,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
911  | 
shyps = Sorts.union sorts shyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
912  | 
hyps = hyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
913  | 
tpairs = tpairs,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
914  | 
prop = Logic.mk_equals  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
915  | 
(Abs (a, T, abstract_over (x, t)), Abs (a, T, abstract_over (x, u)))});  | 
| 21798 | 916  | 
fun check_occs a x ts =  | 
| 
16847
 
8fc160b12e73
invoke_oracle: do not keep theory value, but theory_ref;
 
wenzelm 
parents: 
16725 
diff
changeset
 | 
917  | 
if exists (fn t => Logic.occs (x, t)) ts then  | 
| 21798 | 918  | 
        raise THM ("abstract_rule: variable " ^ quote a ^ " free in assumptions", 0, [th])
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
919  | 
else ();  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
920  | 
in  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
921  | 
case x of  | 
| 21798 | 922  | 
Free (a, _) => (check_occs a x hyps; check_occs a x (terms_of_tpairs tpairs); result)  | 
923  | 
| Var ((a, _), _) => (check_occs a x (terms_of_tpairs tpairs); result)  | 
|
924  | 
    | _ => raise THM ("abstract_rule: not a variable", 0, [th])
 | 
|
| 0 | 925  | 
end;  | 
926  | 
||
927  | 
(*The combination rule  | 
|
| 3529 | 928  | 
f == g t == u  | 
929  | 
--------------  | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
930  | 
f t == g u  | 
| 1220 | 931  | 
*)  | 
| 0 | 932  | 
fun combination th1 th2 =  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
933  | 
let  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
934  | 
    val Thm (der1, {maxidx = max1, shyps = shyps1, hyps = hyps1, tpairs = tpairs1,
 | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
935  | 
prop = prop1, ...}) = th1  | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
936  | 
    and Thm (der2, {maxidx = max2, shyps = shyps2, hyps = hyps2, tpairs = tpairs2,
 | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
937  | 
prop = prop2, ...}) = th2;  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
938  | 
fun chktypes fT tT =  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
939  | 
(case fT of  | 
| 32784 | 940  | 
        Type ("fun", [T1, _]) =>
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
941  | 
if T1 <> tT then  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
942  | 
            raise THM ("combination: types", 0, [th1, th2])
 | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
943  | 
else ()  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
944  | 
      | _ => raise THM ("combination: not function type", 0, [th1, th2]));
 | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
945  | 
in  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
946  | 
case (prop1, prop2) of  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
947  | 
      (Const ("==", Type ("fun", [fT, _])) $ f $ g,
 | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
948  | 
       Const ("==", Type ("fun", [tT, _])) $ t $ u) =>
 | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
949  | 
(chktypes fT tT;  | 
| 37309 | 950  | 
Thm (deriv_rule2 (Proofterm.combination f g t u fT) der1 der2,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
951  | 
           {thy_ref = merge_thys2 th1 th2,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
952  | 
tags = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
953  | 
maxidx = Int.max (max1, max2),  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
954  | 
shyps = Sorts.union shyps1 shyps2,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
955  | 
hyps = union_hyps hyps1 hyps2,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
956  | 
tpairs = union_tpairs tpairs1 tpairs2,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
957  | 
prop = Logic.mk_equals (f $ t, g $ u)}))  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
958  | 
     | _ => raise THM ("combination: premises", 0, [th1, th2])
 | 
| 0 | 959  | 
end;  | 
960  | 
||
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
961  | 
(*Equality introduction  | 
| 3529 | 962  | 
A ==> B B ==> A  | 
963  | 
----------------  | 
|
964  | 
A == B  | 
|
| 1220 | 965  | 
*)  | 
| 0 | 966  | 
fun equal_intr th1 th2 =  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
967  | 
let  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
968  | 
    val Thm (der1, {maxidx = max1, shyps = shyps1, hyps = hyps1, tpairs = tpairs1,
 | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
969  | 
prop = prop1, ...}) = th1  | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
970  | 
    and Thm (der2, {maxidx = max2, shyps = shyps2, hyps = hyps2, tpairs = tpairs2,
 | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
971  | 
prop = prop2, ...}) = th2;  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
972  | 
    fun err msg = raise THM ("equal_intr: " ^ msg, 0, [th1, th2]);
 | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
973  | 
in  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
974  | 
case (prop1, prop2) of  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
975  | 
      (Const("==>", _) $ A $ B, Const("==>", _) $ B' $ A') =>
 | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
976  | 
if A aconv A' andalso B aconv B' then  | 
| 37309 | 977  | 
Thm (deriv_rule2 (Proofterm.equal_intr A B) der1 der2,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
978  | 
           {thy_ref = merge_thys2 th1 th2,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
979  | 
tags = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
980  | 
maxidx = Int.max (max1, max2),  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
981  | 
shyps = Sorts.union shyps1 shyps2,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
982  | 
hyps = union_hyps hyps1 hyps2,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
983  | 
tpairs = union_tpairs tpairs1 tpairs2,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
984  | 
prop = Logic.mk_equals (A, B)})  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
985  | 
else err "not equal"  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
986  | 
| _ => err "premises"  | 
| 1529 | 987  | 
end;  | 
988  | 
||
989  | 
(*The equal propositions rule  | 
|
| 3529 | 990  | 
A == B A  | 
| 1529 | 991  | 
---------  | 
992  | 
B  | 
|
993  | 
*)  | 
|
994  | 
fun equal_elim th1 th2 =  | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
995  | 
let  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
996  | 
    val Thm (der1, {maxidx = max1, shyps = shyps1, hyps = hyps1,
 | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
997  | 
tpairs = tpairs1, prop = prop1, ...}) = th1  | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
998  | 
    and Thm (der2, {maxidx = max2, shyps = shyps2, hyps = hyps2,
 | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
999  | 
tpairs = tpairs2, prop = prop2, ...}) = th2;  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1000  | 
    fun err msg = raise THM ("equal_elim: " ^ msg, 0, [th1, th2]);
 | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1001  | 
in  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1002  | 
case prop1 of  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1003  | 
      Const ("==", _) $ A $ B =>
 | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1004  | 
if prop2 aconv A then  | 
| 37309 | 1005  | 
Thm (deriv_rule2 (Proofterm.equal_elim A B) der1 der2,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1006  | 
           {thy_ref = merge_thys2 th1 th2,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
1007  | 
tags = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1008  | 
maxidx = Int.max (max1, max2),  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1009  | 
shyps = Sorts.union shyps1 shyps2,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1010  | 
hyps = union_hyps hyps1 hyps2,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1011  | 
tpairs = union_tpairs tpairs1 tpairs2,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1012  | 
prop = B})  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1013  | 
else err "not equal"  | 
| 1529 | 1014  | 
| _ => err"major premise"  | 
1015  | 
end;  | 
|
| 0 | 1016  | 
|
| 1220 | 1017  | 
|
1018  | 
||
| 0 | 1019  | 
(**** Derived rules ****)  | 
1020  | 
||
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1021  | 
(*Smash unifies the list of term pairs leaving no flex-flex pairs.  | 
| 
24143
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
1022  | 
Instantiates the theorem and deletes trivial tpairs. Resulting  | 
| 
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
1023  | 
sequence may contain multiple elements if the tpairs are not all  | 
| 
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
1024  | 
flex-flex.*)  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1025  | 
fun flexflex_rule (th as Thm (der, {thy_ref, maxidx, shyps, hyps, tpairs, prop, ...})) =
 | 
| 
24143
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
1026  | 
let val thy = Theory.deref thy_ref in  | 
| 
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
1027  | 
Unify.smash_unifiers thy tpairs (Envir.empty maxidx)  | 
| 
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
1028  | 
|> Seq.map (fn env =>  | 
| 
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
1029  | 
if Envir.is_empty env then th  | 
| 
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
1030  | 
else  | 
| 
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
1031  | 
let  | 
| 
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
1032  | 
val tpairs' = tpairs |> map (pairself (Envir.norm_term env))  | 
| 
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
1033  | 
(*remove trivial tpairs, of the form t==t*)  | 
| 
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
1034  | 
|> filter_out (op aconv);  | 
| 37309 | 1035  | 
val der' = deriv_rule1 (Proofterm.norm_proof' env) der;  | 
| 
24143
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
1036  | 
val prop' = Envir.norm_term env prop;  | 
| 
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
1037  | 
val maxidx = maxidx_tpairs tpairs' (maxidx_of_term prop');  | 
| 
26640
 
92e6d3ec91bd
simplified handling of sorts, removed unnecessary universal witness;
 
wenzelm 
parents: 
26631 
diff
changeset
 | 
1038  | 
val shyps = Envir.insert_sorts env shyps;  | 
| 
24143
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
1039  | 
in  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1040  | 
            Thm (der', {thy_ref = Theory.check_thy thy, tags = [], maxidx = maxidx,
 | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1041  | 
shyps = shyps, hyps = hyps, tpairs = tpairs', prop = prop'})  | 
| 
24143
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
1042  | 
end)  | 
| 
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
1043  | 
end;  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1044  | 
|
| 0 | 1045  | 
|
| 19910 | 1046  | 
(*Generalization of fixed variables  | 
1047  | 
A  | 
|
1048  | 
--------------------  | 
|
1049  | 
A[?'a/'a, ?x/x, ...]  | 
|
1050  | 
*)  | 
|
1051  | 
||
1052  | 
fun generalize ([], []) _ th = th  | 
|
1053  | 
| generalize (tfrees, frees) idx th =  | 
|
1054  | 
let  | 
|
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1055  | 
        val Thm (der, {thy_ref, maxidx, shyps, hyps, tpairs, prop, ...}) = th;
 | 
| 19910 | 1056  | 
        val _ = idx <= maxidx andalso raise THM ("generalize: bad index", idx, [th]);
 | 
1057  | 
||
| 33697 | 1058  | 
val bad_type =  | 
1059  | 
if null tfrees then K false  | 
|
1060  | 
else Term.exists_subtype (fn TFree (a, _) => member (op =) tfrees a | _ => false);  | 
|
| 19910 | 1061  | 
fun bad_term (Free (x, T)) = bad_type T orelse member (op =) frees x  | 
1062  | 
| bad_term (Var (_, T)) = bad_type T  | 
|
1063  | 
| bad_term (Const (_, T)) = bad_type T  | 
|
1064  | 
| bad_term (Abs (_, T, t)) = bad_type T orelse bad_term t  | 
|
1065  | 
| bad_term (t $ u) = bad_term t orelse bad_term u  | 
|
1066  | 
| bad_term (Bound _) = false;  | 
|
1067  | 
val _ = exists bad_term hyps andalso  | 
|
1068  | 
          raise THM ("generalize: variable free in assumptions", 0, [th]);
 | 
|
1069  | 
||
| 31977 | 1070  | 
val gen = Term_Subst.generalize (tfrees, frees) idx;  | 
| 19910 | 1071  | 
val prop' = gen prop;  | 
1072  | 
val tpairs' = map (pairself gen) tpairs;  | 
|
1073  | 
val maxidx' = maxidx_tpairs tpairs' (maxidx_of_term prop');  | 
|
1074  | 
in  | 
|
| 37309 | 1075  | 
Thm (deriv_rule1 (Proofterm.generalize (tfrees, frees) idx) der,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1076  | 
         {thy_ref = thy_ref,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
1077  | 
tags = [],  | 
| 19910 | 1078  | 
maxidx = maxidx',  | 
1079  | 
shyps = shyps,  | 
|
1080  | 
hyps = hyps,  | 
|
1081  | 
tpairs = tpairs',  | 
|
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1082  | 
prop = prop'})  | 
| 19910 | 1083  | 
end;  | 
1084  | 
||
1085  | 
||
| 22584 | 1086  | 
(*Instantiation of schematic variables  | 
| 16656 | 1087  | 
A  | 
1088  | 
--------------------  | 
|
1089  | 
A[t1/v1, ..., tn/vn]  | 
|
| 1220 | 1090  | 
*)  | 
| 0 | 1091  | 
|
| 6928 | 1092  | 
local  | 
1093  | 
||
| 
26939
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26832 
diff
changeset
 | 
1094  | 
fun pretty_typing thy t T = Pretty.block  | 
| 
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26832 
diff
changeset
 | 
1095  | 
[Syntax.pretty_term_global thy t, Pretty.str " ::", Pretty.brk 1, Syntax.pretty_typ_global thy T];  | 
| 15797 | 1096  | 
|
| 
16884
 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 
wenzelm 
parents: 
16847 
diff
changeset
 | 
1097  | 
fun add_inst (ct, cu) (thy_ref, sorts) =  | 
| 6928 | 1098  | 
let  | 
| 
26939
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26832 
diff
changeset
 | 
1099  | 
    val Cterm {t = t, T = T, ...} = ct;
 | 
| 
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26832 
diff
changeset
 | 
1100  | 
    val Cterm {t = u, T = U, sorts = sorts_u, maxidx = maxidx_u, ...} = cu;
 | 
| 
16884
 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 
wenzelm 
parents: 
16847 
diff
changeset
 | 
1101  | 
val thy_ref' = Theory.merge_refs (thy_ref, merge_thys0 ct cu);  | 
| 
 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 
wenzelm 
parents: 
16847 
diff
changeset
 | 
1102  | 
val sorts' = Sorts.union sorts_u sorts;  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1103  | 
in  | 
| 
16884
 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 
wenzelm 
parents: 
16847 
diff
changeset
 | 
1104  | 
(case t of Var v =>  | 
| 20512 | 1105  | 
if T = U then ((v, (u, maxidx_u)), (thy_ref', sorts'))  | 
| 
16884
 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 
wenzelm 
parents: 
16847 
diff
changeset
 | 
1106  | 
else raise TYPE (Pretty.string_of (Pretty.block  | 
| 
 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 
wenzelm 
parents: 
16847 
diff
changeset
 | 
1107  | 
[Pretty.str "instantiate: type conflict",  | 
| 
 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 
wenzelm 
parents: 
16847 
diff
changeset
 | 
1108  | 
Pretty.fbrk, pretty_typing (Theory.deref thy_ref') t T,  | 
| 
 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 
wenzelm 
parents: 
16847 
diff
changeset
 | 
1109  | 
Pretty.fbrk, pretty_typing (Theory.deref thy_ref') u U]), [T, U], [t, u])  | 
| 
 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 
wenzelm 
parents: 
16847 
diff
changeset
 | 
1110  | 
| _ => raise TYPE (Pretty.string_of (Pretty.block  | 
| 
 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 
wenzelm 
parents: 
16847 
diff
changeset
 | 
1111  | 
[Pretty.str "instantiate: not a variable",  | 
| 
26939
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26832 
diff
changeset
 | 
1112  | 
Pretty.fbrk, Syntax.pretty_term_global (Theory.deref thy_ref') t]), [], [t]))  | 
| 0 | 1113  | 
end;  | 
1114  | 
||
| 
16884
 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 
wenzelm 
parents: 
16847 
diff
changeset
 | 
1115  | 
fun add_instT (cT, cU) (thy_ref, sorts) =  | 
| 16656 | 1116  | 
let  | 
| 
16884
 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 
wenzelm 
parents: 
16847 
diff
changeset
 | 
1117  | 
    val Ctyp {T, thy_ref = thy_ref1, ...} = cT
 | 
| 20512 | 1118  | 
    and Ctyp {T = U, thy_ref = thy_ref2, sorts = sorts_U, maxidx = maxidx_U, ...} = cU;
 | 
| 
24143
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
1119  | 
val thy' = Theory.deref (Theory.merge_refs (thy_ref, Theory.merge_refs (thy_ref1, thy_ref2)));  | 
| 
16884
 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 
wenzelm 
parents: 
16847 
diff
changeset
 | 
1120  | 
val sorts' = Sorts.union sorts_U sorts;  | 
| 16656 | 1121  | 
in  | 
| 
16884
 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 
wenzelm 
parents: 
16847 
diff
changeset
 | 
1122  | 
(case T of TVar (v as (_, S)) =>  | 
| 
24143
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
1123  | 
if Sign.of_sort thy' (U, S) then ((v, (U, maxidx_U)), (Theory.check_thy thy', sorts'))  | 
| 
26939
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26832 
diff
changeset
 | 
1124  | 
      else raise TYPE ("Type not of sort " ^ Syntax.string_of_sort_global thy' S, [U], [])
 | 
| 16656 | 1125  | 
| _ => raise TYPE (Pretty.string_of (Pretty.block  | 
| 15797 | 1126  | 
[Pretty.str "instantiate: not a type variable",  | 
| 
26939
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26832 
diff
changeset
 | 
1127  | 
Pretty.fbrk, Syntax.pretty_typ_global thy' T]), [T], []))  | 
| 16656 | 1128  | 
end;  | 
| 0 | 1129  | 
|
| 6928 | 1130  | 
in  | 
1131  | 
||
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1132  | 
(*Left-to-right replacements: ctpairs = [..., (vi, ti), ...].  | 
| 0 | 1133  | 
Instantiates distinct Vars by terms of same type.  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1134  | 
Does NOT normalize the resulting theorem!*)  | 
| 1529 | 1135  | 
fun instantiate ([], []) th = th  | 
| 
16884
 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 
wenzelm 
parents: 
16847 
diff
changeset
 | 
1136  | 
| instantiate (instT, inst) th =  | 
| 16656 | 1137  | 
let  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1138  | 
        val Thm (der, {thy_ref, hyps, shyps, tpairs, prop, ...}) = th;
 | 
| 
16884
 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 
wenzelm 
parents: 
16847 
diff
changeset
 | 
1139  | 
val (inst', (instT', (thy_ref', shyps'))) =  | 
| 
 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 
wenzelm 
parents: 
16847 
diff
changeset
 | 
1140  | 
(thy_ref, shyps) |> fold_map add_inst inst ||> fold_map add_instT instT;  | 
| 31977 | 1141  | 
val subst = Term_Subst.instantiate_maxidx (instT', inst');  | 
| 20512 | 1142  | 
val (prop', maxidx1) = subst prop ~1;  | 
1143  | 
val (tpairs', maxidx') =  | 
|
1144  | 
fold_map (fn (t, u) => fn i => subst t i ||>> subst u) tpairs maxidx1;  | 
|
| 16656 | 1145  | 
in  | 
| 37309 | 1146  | 
Thm (deriv_rule1  | 
1147  | 
(fn d => Proofterm.instantiate (map (apsnd #1) instT', map (apsnd #1) inst') d) der,  | 
|
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1148  | 
         {thy_ref = thy_ref',
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
1149  | 
tags = [],  | 
| 
20545
 
4c1068697159
instantiate: omit has_duplicates check, which is irrelevant for soundness;
 
wenzelm 
parents: 
20512 
diff
changeset
 | 
1150  | 
maxidx = maxidx',  | 
| 
 
4c1068697159
instantiate: omit has_duplicates check, which is irrelevant for soundness;
 
wenzelm 
parents: 
20512 
diff
changeset
 | 
1151  | 
shyps = shyps',  | 
| 
 
4c1068697159
instantiate: omit has_duplicates check, which is irrelevant for soundness;
 
wenzelm 
parents: 
20512 
diff
changeset
 | 
1152  | 
hyps = hyps,  | 
| 
 
4c1068697159
instantiate: omit has_duplicates check, which is irrelevant for soundness;
 
wenzelm 
parents: 
20512 
diff
changeset
 | 
1153  | 
tpairs = tpairs',  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1154  | 
prop = prop'})  | 
| 16656 | 1155  | 
end  | 
1156  | 
handle TYPE (msg, _, _) => raise THM (msg, 0, [th]);  | 
|
| 6928 | 1157  | 
|
| 22584 | 1158  | 
fun instantiate_cterm ([], []) ct = ct  | 
1159  | 
| instantiate_cterm (instT, inst) ct =  | 
|
1160  | 
let  | 
|
1161  | 
        val Cterm {thy_ref, t, T, sorts, ...} = ct;
 | 
|
1162  | 
val (inst', (instT', (thy_ref', sorts'))) =  | 
|
1163  | 
(thy_ref, sorts) |> fold_map add_inst inst ||> fold_map add_instT instT;  | 
|
| 31977 | 1164  | 
val subst = Term_Subst.instantiate_maxidx (instT', inst');  | 
1165  | 
val substT = Term_Subst.instantiateT_maxidx instT';  | 
|
| 22584 | 1166  | 
val (t', maxidx1) = subst t ~1;  | 
1167  | 
val (T', maxidx') = substT T maxidx1;  | 
|
1168  | 
      in Cterm {thy_ref = thy_ref', t = t', T = T', sorts = sorts', maxidx = maxidx'} end
 | 
|
1169  | 
handle TYPE (msg, _, _) => raise CTERM (msg, [ct]);  | 
|
1170  | 
||
| 6928 | 1171  | 
end;  | 
1172  | 
||
| 0 | 1173  | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1174  | 
(*The trivial implication A ==> A, justified by assume and forall rules.  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1175  | 
A can contain Vars, not so for assume!*)  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1176  | 
fun trivial (Cterm {thy_ref, t =A, T, maxidx, sorts}) =
 | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1177  | 
if T <> propT then  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1178  | 
    raise THM ("trivial: the term must have type prop", 0, [])
 | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1179  | 
else  | 
| 37309 | 1180  | 
    Thm (deriv_rule0 (Proofterm.AbsP ("H", NONE, Proofterm.PBound 0)),
 | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1181  | 
     {thy_ref = thy_ref,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
1182  | 
tags = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1183  | 
maxidx = maxidx,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1184  | 
shyps = sorts,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1185  | 
hyps = [],  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1186  | 
tpairs = [],  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1187  | 
prop = Logic.mk_implies (A, A)});  | 
| 0 | 1188  | 
|
| 31944 | 1189  | 
(*Axiom-scheme reflecting signature contents  | 
1190  | 
T :: c  | 
|
1191  | 
-------------------  | 
|
1192  | 
OFCLASS(T, c_class)  | 
|
1193  | 
*)  | 
|
1194  | 
fun of_class (cT, raw_c) =  | 
|
| 
24143
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
1195  | 
let  | 
| 31944 | 1196  | 
    val Ctyp {thy_ref, T, ...} = cT;
 | 
1197  | 
val thy = Theory.deref thy_ref;  | 
|
| 31903 | 1198  | 
val c = Sign.certify_class thy raw_c;  | 
| 31944 | 1199  | 
    val Cterm {t = prop, maxidx, sorts, ...} = cterm_of thy (Logic.mk_of_class (T, c));
 | 
| 
399
 
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
 
wenzelm 
parents: 
387 
diff
changeset
 | 
1200  | 
in  | 
| 31944 | 1201  | 
if Sign.of_sort thy (T, [c]) then  | 
| 37309 | 1202  | 
Thm (deriv_rule0 (Proofterm.OfClass (T, c)),  | 
| 31944 | 1203  | 
       {thy_ref = Theory.check_thy thy,
 | 
1204  | 
tags = [],  | 
|
1205  | 
maxidx = maxidx,  | 
|
1206  | 
shyps = sorts,  | 
|
1207  | 
hyps = [],  | 
|
1208  | 
tpairs = [],  | 
|
1209  | 
prop = prop})  | 
|
1210  | 
    else raise THM ("of_class: type not of class " ^ Syntax.string_of_sort_global thy [c], 0, [])
 | 
|
| 
399
 
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
 
wenzelm 
parents: 
387 
diff
changeset
 | 
1211  | 
end;  | 
| 
 
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
 
wenzelm 
parents: 
387 
diff
changeset
 | 
1212  | 
|
| 
36614
 
b6c031ad3690
minor tuning of Thm.strip_shyps -- no longer pervasive;
 
wenzelm 
parents: 
36613 
diff
changeset
 | 
1213  | 
(*Remove extra sorts that are witnessed by type signature information*)  | 
| 
 
b6c031ad3690
minor tuning of Thm.strip_shyps -- no longer pervasive;
 
wenzelm 
parents: 
36613 
diff
changeset
 | 
1214  | 
fun strip_shyps (thm as Thm (_, {shyps = [], ...})) = thm
 | 
| 
 
b6c031ad3690
minor tuning of Thm.strip_shyps -- no longer pervasive;
 
wenzelm 
parents: 
36613 
diff
changeset
 | 
1215  | 
  | strip_shyps (thm as Thm (der, {thy_ref, tags, maxidx, shyps, hyps, tpairs, prop})) =
 | 
| 
 
b6c031ad3690
minor tuning of Thm.strip_shyps -- no longer pervasive;
 
wenzelm 
parents: 
36613 
diff
changeset
 | 
1216  | 
let  | 
| 
 
b6c031ad3690
minor tuning of Thm.strip_shyps -- no longer pervasive;
 
wenzelm 
parents: 
36613 
diff
changeset
 | 
1217  | 
val thy = Theory.deref thy_ref;  | 
| 
36621
 
2fd4e2c76636
proof terms for strip_shyps, based on the version by krauss/schropp with some notable differences:
 
wenzelm 
parents: 
36619 
diff
changeset
 | 
1218  | 
val algebra = Sign.classes_of thy;  | 
| 
 
2fd4e2c76636
proof terms for strip_shyps, based on the version by krauss/schropp with some notable differences:
 
wenzelm 
parents: 
36619 
diff
changeset
 | 
1219  | 
|
| 
 
2fd4e2c76636
proof terms for strip_shyps, based on the version by krauss/schropp with some notable differences:
 
wenzelm 
parents: 
36619 
diff
changeset
 | 
1220  | 
val present = (fold_terms o fold_types o fold_atyps_sorts) (insert (eq_fst op =)) thm [];  | 
| 
36614
 
b6c031ad3690
minor tuning of Thm.strip_shyps -- no longer pervasive;
 
wenzelm 
parents: 
36613 
diff
changeset
 | 
1221  | 
val extra = fold (Sorts.remove_sort o #2) present shyps;  | 
| 
 
b6c031ad3690
minor tuning of Thm.strip_shyps -- no longer pervasive;
 
wenzelm 
parents: 
36613 
diff
changeset
 | 
1222  | 
val witnessed = Sign.witness_sorts thy present extra;  | 
| 
 
b6c031ad3690
minor tuning of Thm.strip_shyps -- no longer pervasive;
 
wenzelm 
parents: 
36613 
diff
changeset
 | 
1223  | 
val extra' = fold (Sorts.remove_sort o #2) witnessed extra  | 
| 
36621
 
2fd4e2c76636
proof terms for strip_shyps, based on the version by krauss/schropp with some notable differences:
 
wenzelm 
parents: 
36619 
diff
changeset
 | 
1224  | 
|> Sorts.minimal_sorts algebra;  | 
| 
36614
 
b6c031ad3690
minor tuning of Thm.strip_shyps -- no longer pervasive;
 
wenzelm 
parents: 
36613 
diff
changeset
 | 
1225  | 
val shyps' = fold (Sorts.insert_sort o #2) present extra';  | 
| 
 
b6c031ad3690
minor tuning of Thm.strip_shyps -- no longer pervasive;
 
wenzelm 
parents: 
36613 
diff
changeset
 | 
1226  | 
in  | 
| 37309 | 1227  | 
Thm (deriv_rule_unconditional  | 
1228  | 
(Proofterm.strip_shyps_proof algebra present witnessed extra') der,  | 
|
| 
36621
 
2fd4e2c76636
proof terms for strip_shyps, based on the version by krauss/schropp with some notable differences:
 
wenzelm 
parents: 
36619 
diff
changeset
 | 
1229  | 
         {thy_ref = Theory.check_thy thy, tags = tags, maxidx = maxidx,
 | 
| 
36614
 
b6c031ad3690
minor tuning of Thm.strip_shyps -- no longer pervasive;
 
wenzelm 
parents: 
36613 
diff
changeset
 | 
1230  | 
shyps = shyps', hyps = hyps, tpairs = tpairs, prop = prop})  | 
| 
 
b6c031ad3690
minor tuning of Thm.strip_shyps -- no longer pervasive;
 
wenzelm 
parents: 
36613 
diff
changeset
 | 
1231  | 
end;  | 
| 
 
b6c031ad3690
minor tuning of Thm.strip_shyps -- no longer pervasive;
 
wenzelm 
parents: 
36613 
diff
changeset
 | 
1232  | 
|
| 
36769
 
b6b88bf695b3
Thm.unconstrainT based on Logic.unconstrainT -- no proofterm yet;
 
wenzelm 
parents: 
36768 
diff
changeset
 | 
1233  | 
(*Internalize sort constraints of type variables*)  | 
| 
36883
 
4ed0d72def50
added Proofterm.unconstrain_thm_proof for Thm.unconstrainT -- loosely based on the version by krauss/schropp;
 
wenzelm 
parents: 
36882 
diff
changeset
 | 
1234  | 
fun unconstrainT (thm as Thm (der, args)) =  | 
| 19505 | 1235  | 
let  | 
| 
36883
 
4ed0d72def50
added Proofterm.unconstrain_thm_proof for Thm.unconstrainT -- loosely based on the version by krauss/schropp;
 
wenzelm 
parents: 
36882 
diff
changeset
 | 
1236  | 
    val Deriv {promises, body} = der;
 | 
| 
 
4ed0d72def50
added Proofterm.unconstrain_thm_proof for Thm.unconstrainT -- loosely based on the version by krauss/schropp;
 
wenzelm 
parents: 
36882 
diff
changeset
 | 
1237  | 
    val {thy_ref, shyps, hyps, tpairs, prop, ...} = args;
 | 
| 
 
4ed0d72def50
added Proofterm.unconstrain_thm_proof for Thm.unconstrainT -- loosely based on the version by krauss/schropp;
 
wenzelm 
parents: 
36882 
diff
changeset
 | 
1238  | 
|
| 
36769
 
b6b88bf695b3
Thm.unconstrainT based on Logic.unconstrainT -- no proofterm yet;
 
wenzelm 
parents: 
36768 
diff
changeset
 | 
1239  | 
    fun err msg = raise THM ("unconstrainT: " ^ msg, 0, [thm]);
 | 
| 
 
b6b88bf695b3
Thm.unconstrainT based on Logic.unconstrainT -- no proofterm yet;
 
wenzelm 
parents: 
36768 
diff
changeset
 | 
1240  | 
val _ = null hyps orelse err "illegal hyps";  | 
| 
 
b6b88bf695b3
Thm.unconstrainT based on Logic.unconstrainT -- no proofterm yet;
 
wenzelm 
parents: 
36768 
diff
changeset
 | 
1241  | 
val _ = null tpairs orelse err "unsolved flex-flex constraints";  | 
| 
 
b6b88bf695b3
Thm.unconstrainT based on Logic.unconstrainT -- no proofterm yet;
 
wenzelm 
parents: 
36768 
diff
changeset
 | 
1242  | 
val tfrees = rev (Term.add_tfree_names prop []);  | 
| 
 
b6b88bf695b3
Thm.unconstrainT based on Logic.unconstrainT -- no proofterm yet;
 
wenzelm 
parents: 
36768 
diff
changeset
 | 
1243  | 
    val _ = null tfrees orelse err ("illegal free type variables " ^ commas_quote tfrees);
 | 
| 
 
b6b88bf695b3
Thm.unconstrainT based on Logic.unconstrainT -- no proofterm yet;
 
wenzelm 
parents: 
36768 
diff
changeset
 | 
1244  | 
|
| 
41700
 
f33d5a00c25d
thm_proof: visible fulfill_body only, without joining nested thms -- retain proof irrelevance, which is important for parallel performance;
 
wenzelm 
parents: 
41699 
diff
changeset
 | 
1245  | 
val ps = map (apsnd (Future.map fulfill_body)) promises;  | 
| 
36883
 
4ed0d72def50
added Proofterm.unconstrain_thm_proof for Thm.unconstrainT -- loosely based on the version by krauss/schropp;
 
wenzelm 
parents: 
36882 
diff
changeset
 | 
1246  | 
val thy = Theory.deref thy_ref;  | 
| 37309 | 1247  | 
val (pthm as (_, (_, prop', _)), proof) =  | 
1248  | 
Proofterm.unconstrain_thm_proof thy shyps prop ps body;  | 
|
| 
36883
 
4ed0d72def50
added Proofterm.unconstrain_thm_proof for Thm.unconstrainT -- loosely based on the version by krauss/schropp;
 
wenzelm 
parents: 
36882 
diff
changeset
 | 
1249  | 
val der' = make_deriv [] [] [pthm] proof;  | 
| 
 
4ed0d72def50
added Proofterm.unconstrain_thm_proof for Thm.unconstrainT -- loosely based on the version by krauss/schropp;
 
wenzelm 
parents: 
36882 
diff
changeset
 | 
1250  | 
val _ = Theory.check_thy thy;  | 
| 19505 | 1251  | 
in  | 
| 
36883
 
4ed0d72def50
added Proofterm.unconstrain_thm_proof for Thm.unconstrainT -- loosely based on the version by krauss/schropp;
 
wenzelm 
parents: 
36882 
diff
changeset
 | 
1252  | 
Thm (der',  | 
| 
36769
 
b6b88bf695b3
Thm.unconstrainT based on Logic.unconstrainT -- no proofterm yet;
 
wenzelm 
parents: 
36768 
diff
changeset
 | 
1253  | 
     {thy_ref = thy_ref,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
1254  | 
tags = [],  | 
| 
36769
 
b6b88bf695b3
Thm.unconstrainT based on Logic.unconstrainT -- no proofterm yet;
 
wenzelm 
parents: 
36768 
diff
changeset
 | 
1255  | 
maxidx = maxidx_of_term prop',  | 
| 
 
b6b88bf695b3
Thm.unconstrainT based on Logic.unconstrainT -- no proofterm yet;
 
wenzelm 
parents: 
36768 
diff
changeset
 | 
1256  | 
shyps = [[]], (*potentially redundant*)  | 
| 
36883
 
4ed0d72def50
added Proofterm.unconstrain_thm_proof for Thm.unconstrainT -- loosely based on the version by krauss/schropp;
 
wenzelm 
parents: 
36882 
diff
changeset
 | 
1257  | 
hyps = [],  | 
| 
 
4ed0d72def50
added Proofterm.unconstrain_thm_proof for Thm.unconstrainT -- loosely based on the version by krauss/schropp;
 
wenzelm 
parents: 
36882 
diff
changeset
 | 
1258  | 
tpairs = [],  | 
| 
36769
 
b6b88bf695b3
Thm.unconstrainT based on Logic.unconstrainT -- no proofterm yet;
 
wenzelm 
parents: 
36768 
diff
changeset
 | 
1259  | 
prop = prop'})  | 
| 19505 | 1260  | 
end;  | 
| 
399
 
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
 
wenzelm 
parents: 
387 
diff
changeset
 | 
1261  | 
|
| 6786 | 1262  | 
(* Replace all TFrees not fixed or in the hyps by new TVars *)  | 
| 
35845
 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 
wenzelm 
parents: 
35408 
diff
changeset
 | 
1263  | 
fun varifyT_global' fixed (Thm (der, {thy_ref, maxidx, shyps, hyps, tpairs, prop, ...})) =
 | 
| 12500 | 1264  | 
let  | 
| 29272 | 1265  | 
val tfrees = fold Term.add_tfrees hyps fixed;  | 
| 
13658
 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 
berghofe 
parents: 
13642 
diff
changeset
 | 
1266  | 
val prop1 = attach_tpairs tpairs prop;  | 
| 
35845
 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 
wenzelm 
parents: 
35408 
diff
changeset
 | 
1267  | 
val (al, prop2) = Type.varify_global tfrees prop1;  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1268  | 
val (ts, prop3) = Logic.strip_prems (length tpairs, [], prop2);  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1269  | 
in  | 
| 37309 | 1270  | 
(al, Thm (deriv_rule1 (Proofterm.varify_proof prop tfrees) der,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1271  | 
     {thy_ref = thy_ref,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
1272  | 
tags = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1273  | 
maxidx = Int.max (0, maxidx),  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1274  | 
shyps = shyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1275  | 
hyps = hyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1276  | 
tpairs = rev (map Logic.dest_equals ts),  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1277  | 
prop = prop3}))  | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1278  | 
end;  | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1279  | 
|
| 
35845
 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 
wenzelm 
parents: 
35408 
diff
changeset
 | 
1280  | 
val varifyT_global = #2 o varifyT_global' [];  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1281  | 
|
| 
36615
 
88756a5a92fc
renamed Thm.freezeT to Thm.legacy_freezeT -- it is based on Type.legacy_freeze;
 
wenzelm 
parents: 
36614 
diff
changeset
 | 
1282  | 
(* Replace all TVars by TFrees that are often new *)  | 
| 
 
88756a5a92fc
renamed Thm.freezeT to Thm.legacy_freezeT -- it is based on Type.legacy_freeze;
 
wenzelm 
parents: 
36614 
diff
changeset
 | 
1283  | 
fun legacy_freezeT (Thm (der, {thy_ref, shyps, hyps, tpairs, prop, ...})) =
 | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1284  | 
let  | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1285  | 
val prop1 = attach_tpairs tpairs prop;  | 
| 33832 | 1286  | 
val prop2 = Type.legacy_freeze prop1;  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1287  | 
val (ts, prop3) = Logic.strip_prems (length tpairs, [], prop2);  | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1288  | 
in  | 
| 37309 | 1289  | 
Thm (deriv_rule1 (Proofterm.legacy_freezeT prop1) der,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1290  | 
     {thy_ref = thy_ref,
 | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1291  | 
tags = [],  | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1292  | 
maxidx = maxidx_of_term prop2,  | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1293  | 
shyps = shyps,  | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1294  | 
hyps = hyps,  | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1295  | 
tpairs = rev (map Logic.dest_equals ts),  | 
| 18127 | 1296  | 
prop = prop3})  | 
| 0 | 1297  | 
end;  | 
1298  | 
||
1299  | 
||
1300  | 
(*** Inference rules for tactics ***)  | 
|
1301  | 
||
1302  | 
(*Destruct proof state into constraints, other goals, goal(i), rest *)  | 
|
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1303  | 
fun dest_state (state as Thm (_, {prop,tpairs,...}), i) =
 | 
| 
13658
 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 
berghofe 
parents: 
13642 
diff
changeset
 | 
1304  | 
(case Logic.strip_prems(i, [], prop) of  | 
| 
 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 
berghofe 
parents: 
13642 
diff
changeset
 | 
1305  | 
(B::rBs, C) => (tpairs, rev rBs, B, C)  | 
| 
 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 
berghofe 
parents: 
13642 
diff
changeset
 | 
1306  | 
    | _ => raise THM("dest_state", i, [state]))
 | 
| 0 | 1307  | 
  handle TERM _ => raise THM("dest_state", i, [state]);
 | 
1308  | 
||
| 309 | 1309  | 
(*Increment variables and parameters of orule as required for  | 
| 18035 | 1310  | 
resolution with a goal.*)  | 
1311  | 
fun lift_rule goal orule =  | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1312  | 
let  | 
| 18035 | 1313  | 
    val Cterm {t = gprop, T, maxidx = gmax, sorts, ...} = goal;
 | 
1314  | 
val inc = gmax + 1;  | 
|
1315  | 
val lift_abs = Logic.lift_abs inc gprop;  | 
|
1316  | 
val lift_all = Logic.lift_all inc gprop;  | 
|
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1317  | 
    val Thm (der, {maxidx, shyps, hyps, tpairs, prop, ...}) = orule;
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1318  | 
val (As, B) = Logic.strip_horn prop;  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1319  | 
in  | 
| 18035 | 1320  | 
    if T <> propT then raise THM ("lift_rule: the term must have type prop", 0, [])
 | 
1321  | 
else  | 
|
| 37309 | 1322  | 
Thm (deriv_rule1 (Proofterm.lift_proof gprop inc prop) der,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1323  | 
       {thy_ref = merge_thys1 goal orule,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
1324  | 
tags = [],  | 
| 18035 | 1325  | 
maxidx = maxidx + inc,  | 
1326  | 
shyps = Sorts.union shyps sorts, (*sic!*)  | 
|
1327  | 
hyps = hyps,  | 
|
1328  | 
tpairs = map (pairself lift_abs) tpairs,  | 
|
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1329  | 
prop = Logic.list_implies (map lift_all As, lift_all B)})  | 
| 0 | 1330  | 
end;  | 
1331  | 
||
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1332  | 
fun incr_indexes i (thm as Thm (der, {thy_ref, maxidx, shyps, hyps, tpairs, prop, ...})) =
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1333  | 
  if i < 0 then raise THM ("negative increment", 0, [thm])
 | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1334  | 
else if i = 0 then thm  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1335  | 
else  | 
| 37309 | 1336  | 
Thm (deriv_rule1 (Proofterm.incr_indexes i) der,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1337  | 
     {thy_ref = thy_ref,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
1338  | 
tags = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1339  | 
maxidx = maxidx + i,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1340  | 
shyps = shyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1341  | 
hyps = hyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1342  | 
tpairs = map (pairself (Logic.incr_indexes ([], i))) tpairs,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1343  | 
prop = Logic.incr_indexes ([], i) prop});  | 
| 
10416
 
5b33e732e459
- Moved rewriting functions to meta_simplifier.ML
 
berghofe 
parents: 
10348 
diff
changeset
 | 
1344  | 
|
| 0 | 1345  | 
(*Solve subgoal Bi of proof state B1...Bn/C by assumption. *)  | 
1346  | 
fun assumption i state =  | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1347  | 
let  | 
| 32784 | 1348  | 
    val Thm (der, {thy_ref, maxidx, shyps, hyps, ...}) = state;
 | 
| 16656 | 1349  | 
val thy = Theory.deref thy_ref;  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1350  | 
val (tpairs, Bs, Bi, C) = dest_state (state, i);  | 
| 32032 | 1351  | 
fun newth n (env, tpairs) =  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1352  | 
Thm (deriv_rule1  | 
| 37309 | 1353  | 
((if Envir.is_empty env then I else (Proofterm.norm_proof' env)) o  | 
1354  | 
Proofterm.assumption_proof Bs Bi n) der,  | 
|
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1355  | 
       {tags = [],
 | 
| 32032 | 1356  | 
maxidx = Envir.maxidx_of env,  | 
| 
26640
 
92e6d3ec91bd
simplified handling of sorts, removed unnecessary universal witness;
 
wenzelm 
parents: 
26631 
diff
changeset
 | 
1357  | 
shyps = Envir.insert_sorts env shyps,  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1358  | 
hyps = hyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1359  | 
tpairs =  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1360  | 
if Envir.is_empty env then tpairs  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1361  | 
else map (pairself (Envir.norm_term env)) tpairs,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1362  | 
prop =  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1363  | 
if Envir.is_empty env then (*avoid wasted normalizations*)  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1364  | 
Logic.list_implies (Bs, C)  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1365  | 
else (*normalize the new rule fully*)  | 
| 
24143
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
1366  | 
Envir.norm_term env (Logic.list_implies (Bs, C)),  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1367  | 
thy_ref = Theory.check_thy thy});  | 
| 
30554
 
73f8bd5f0af8
substantial speedup of assumption and elim-resolution: Logic.assum_problems refrains from eager application of parameters, discriminate via Term.could_unify before invoking full unification (assumes terms in beta-normal form, as do higher term net operations anyway);
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
1368  | 
|
| 
30556
 
7be15917f3fa
eq_assumption: slightly more efficient by checking (open) result of Logic.assum_problems directly;
 
wenzelm 
parents: 
30554 
diff
changeset
 | 
1369  | 
val (close, asms, concl) = Logic.assum_problems (~1, Bi);  | 
| 
 
7be15917f3fa
eq_assumption: slightly more efficient by checking (open) result of Logic.assum_problems directly;
 
wenzelm 
parents: 
30554 
diff
changeset
 | 
1370  | 
val concl' = close concl;  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1371  | 
fun addprfs [] _ = Seq.empty  | 
| 
30556
 
7be15917f3fa
eq_assumption: slightly more efficient by checking (open) result of Logic.assum_problems directly;
 
wenzelm 
parents: 
30554 
diff
changeset
 | 
1372  | 
| addprfs (asm :: rest) n = Seq.make (fn () => Seq.pull  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1373  | 
(Seq.mapp (newth n)  | 
| 
30556
 
7be15917f3fa
eq_assumption: slightly more efficient by checking (open) result of Logic.assum_problems directly;
 
wenzelm 
parents: 
30554 
diff
changeset
 | 
1374  | 
(if Term.could_unify (asm, concl) then  | 
| 
 
7be15917f3fa
eq_assumption: slightly more efficient by checking (open) result of Logic.assum_problems directly;
 
wenzelm 
parents: 
30554 
diff
changeset
 | 
1375  | 
(Unify.unifiers (thy, Envir.empty maxidx, (close asm, concl') :: tpairs))  | 
| 
30554
 
73f8bd5f0af8
substantial speedup of assumption and elim-resolution: Logic.assum_problems refrains from eager application of parameters, discriminate via Term.could_unify before invoking full unification (assumes terms in beta-normal form, as do higher term net operations anyway);
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
1376  | 
else Seq.empty)  | 
| 
 
73f8bd5f0af8
substantial speedup of assumption and elim-resolution: Logic.assum_problems refrains from eager application of parameters, discriminate via Term.could_unify before invoking full unification (assumes terms in beta-normal form, as do higher term net operations anyway);
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
1377  | 
(addprfs rest (n + 1))))  | 
| 
30556
 
7be15917f3fa
eq_assumption: slightly more efficient by checking (open) result of Logic.assum_problems directly;
 
wenzelm 
parents: 
30554 
diff
changeset
 | 
1378  | 
in addprfs asms 1 end;  | 
| 0 | 1379  | 
|
| 250 | 1380  | 
(*Solve subgoal Bi of proof state B1...Bn/C by assumption.  | 
| 0 | 1381  | 
Checks if Bi's conclusion is alpha-convertible to one of its assumptions*)  | 
1382  | 
fun eq_assumption i state =  | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1383  | 
let  | 
| 32784 | 1384  | 
    val Thm (der, {thy_ref, maxidx, shyps, hyps, ...}) = state;
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1385  | 
val (tpairs, Bs, Bi, C) = dest_state (state, i);  | 
| 
30556
 
7be15917f3fa
eq_assumption: slightly more efficient by checking (open) result of Logic.assum_problems directly;
 
wenzelm 
parents: 
30554 
diff
changeset
 | 
1386  | 
val (_, asms, concl) = Logic.assum_problems (~1, Bi);  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1387  | 
in  | 
| 
30556
 
7be15917f3fa
eq_assumption: slightly more efficient by checking (open) result of Logic.assum_problems directly;
 
wenzelm 
parents: 
30554 
diff
changeset
 | 
1388  | 
(case find_index (fn asm => Pattern.aeconv (asm, concl)) asms of  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1389  | 
      ~1 => raise THM ("eq_assumption", 0, [state])
 | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1390  | 
| n =>  | 
| 37309 | 1391  | 
Thm (deriv_rule1 (Proofterm.assumption_proof Bs Bi (n + 1)) der,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1392  | 
         {thy_ref = thy_ref,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
1393  | 
tags = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1394  | 
maxidx = maxidx,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1395  | 
shyps = shyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1396  | 
hyps = hyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1397  | 
tpairs = tpairs,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1398  | 
prop = Logic.list_implies (Bs, C)}))  | 
| 0 | 1399  | 
end;  | 
1400  | 
||
1401  | 
||
| 2671 | 1402  | 
(*For rotate_tac: fast rotation of assumptions of subgoal i*)  | 
1403  | 
fun rotate_rule k i state =  | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1404  | 
let  | 
| 32784 | 1405  | 
    val Thm (der, {thy_ref, maxidx, shyps, hyps, ...}) = state;
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1406  | 
val (tpairs, Bs, Bi, C) = dest_state (state, i);  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1407  | 
val params = Term.strip_all_vars Bi  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1408  | 
and rest = Term.strip_all_body Bi;  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1409  | 
val asms = Logic.strip_imp_prems rest  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1410  | 
and concl = Logic.strip_imp_concl rest;  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1411  | 
val n = length asms;  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1412  | 
val m = if k < 0 then n + k else k;  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1413  | 
val Bi' =  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1414  | 
if 0 = m orelse m = n then Bi  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1415  | 
else if 0 < m andalso m < n then  | 
| 19012 | 1416  | 
let val (ps, qs) = chop m asms  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1417  | 
in list_all (params, Logic.list_implies (qs @ ps, concl)) end  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1418  | 
      else raise THM ("rotate_rule", k, [state]);
 | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1419  | 
in  | 
| 37309 | 1420  | 
Thm (deriv_rule1 (Proofterm.rotate_proof Bs Bi m) der,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1421  | 
     {thy_ref = thy_ref,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
1422  | 
tags = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1423  | 
maxidx = maxidx,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1424  | 
shyps = shyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1425  | 
hyps = hyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1426  | 
tpairs = tpairs,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1427  | 
prop = Logic.list_implies (Bs @ [Bi'], C)})  | 
| 2671 | 1428  | 
end;  | 
1429  | 
||
1430  | 
||
| 
7248
 
322151fe6f02
new primitive rule permute_prems to underlie defer_tac and rotate_prems
 
paulson 
parents: 
7070 
diff
changeset
 | 
1431  | 
(*Rotates a rule's premises to the left by k, leaving the first j premises  | 
| 
 
322151fe6f02
new primitive rule permute_prems to underlie defer_tac and rotate_prems
 
paulson 
parents: 
7070 
diff
changeset
 | 
1432  | 
unchanged. Does nothing if k=0 or if k equals n-j, where n is the  | 
| 16656 | 1433  | 
number of premises. Useful with etac and underlies defer_tac*)  | 
| 
7248
 
322151fe6f02
new primitive rule permute_prems to underlie defer_tac and rotate_prems
 
paulson 
parents: 
7070 
diff
changeset
 | 
1434  | 
fun permute_prems j k rl =  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1435  | 
let  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1436  | 
    val Thm (der, {thy_ref, maxidx, shyps, hyps, tpairs, prop, ...}) = rl;
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1437  | 
val prems = Logic.strip_imp_prems prop  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1438  | 
and concl = Logic.strip_imp_concl prop;  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1439  | 
val moved_prems = List.drop (prems, j)  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1440  | 
and fixed_prems = List.take (prems, j)  | 
| 43278 | 1441  | 
      handle General.Subscript => raise THM ("permute_prems: j", j, [rl]);
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1442  | 
val n_j = length moved_prems;  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1443  | 
val m = if k < 0 then n_j + k else k;  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1444  | 
val prop' =  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1445  | 
if 0 = m orelse m = n_j then prop  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1446  | 
else if 0 < m andalso m < n_j then  | 
| 19012 | 1447  | 
let val (ps, qs) = chop m moved_prems  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1448  | 
in Logic.list_implies (fixed_prems @ qs @ ps, concl) end  | 
| 16725 | 1449  | 
      else raise THM ("permute_prems: k", k, [rl]);
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1450  | 
in  | 
| 37309 | 1451  | 
Thm (deriv_rule1 (Proofterm.permute_prems_proof prems j m) der,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1452  | 
     {thy_ref = thy_ref,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
1453  | 
tags = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1454  | 
maxidx = maxidx,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1455  | 
shyps = shyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1456  | 
hyps = hyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1457  | 
tpairs = tpairs,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1458  | 
prop = prop'})  | 
| 
7248
 
322151fe6f02
new primitive rule permute_prems to underlie defer_tac and rotate_prems
 
paulson 
parents: 
7070 
diff
changeset
 | 
1459  | 
end;  | 
| 
 
322151fe6f02
new primitive rule permute_prems to underlie defer_tac and rotate_prems
 
paulson 
parents: 
7070 
diff
changeset
 | 
1460  | 
|
| 
 
322151fe6f02
new primitive rule permute_prems to underlie defer_tac and rotate_prems
 
paulson 
parents: 
7070 
diff
changeset
 | 
1461  | 
|
| 0 | 1462  | 
(** User renaming of parameters in a subgoal **)  | 
1463  | 
||
1464  | 
(*Calls error rather than raising an exception because it is intended  | 
|
1465  | 
for top-level use -- exception handling would not make sense here.  | 
|
1466  | 
The names in cs, if distinct, are used for the innermost parameters;  | 
|
| 17868 | 1467  | 
preceding parameters may be renamed to make all params distinct.*)  | 
| 0 | 1468  | 
fun rename_params_rule (cs, i) state =  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1469  | 
let  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1470  | 
    val Thm (der, {thy_ref, tags, maxidx, shyps, hyps, ...}) = state;
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1471  | 
val (tpairs, Bs, Bi, C) = dest_state (state, i);  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1472  | 
val iparams = map #1 (Logic.strip_params Bi);  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1473  | 
val short = length iparams - length cs;  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1474  | 
val newnames =  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1475  | 
if short < 0 then error "More names than abstractions!"  | 
| 33957 | 1476  | 
else Name.variant_list cs (take short iparams) @ cs;  | 
| 20330 | 1477  | 
val freenames = Term.fold_aterms (fn Free (x, _) => insert (op =) x | _ => I) Bi [];  | 
| 45328 | 1478  | 
val newBi = Logic.list_rename_params newnames Bi;  | 
| 250 | 1479  | 
in  | 
| 21182 | 1480  | 
(case duplicates (op =) cs of  | 
1481  | 
      a :: _ => (warning ("Can't rename.  Bound variables not distinct: " ^ a); state)
 | 
|
1482  | 
| [] =>  | 
|
| 
33049
 
c38f02fdf35d
curried inter as canonical list operation (beware of argument order)
 
haftmann 
parents: 
33038 
diff
changeset
 | 
1483  | 
(case inter (op =) cs freenames of  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1484  | 
        a :: _ => (warning ("Can't rename.  Bound/Free variable clash: " ^ a); state)
 | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1485  | 
| [] =>  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1486  | 
Thm (der,  | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1487  | 
         {thy_ref = thy_ref,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
1488  | 
tags = tags,  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1489  | 
maxidx = maxidx,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1490  | 
shyps = shyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1491  | 
hyps = hyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1492  | 
tpairs = tpairs,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1493  | 
prop = Logic.list_implies (Bs @ [newBi], C)})))  | 
| 0 | 1494  | 
end;  | 
1495  | 
||
| 12982 | 1496  | 
|
| 0 | 1497  | 
(*** Preservation of bound variable names ***)  | 
1498  | 
||
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1499  | 
fun rename_boundvars pat obj (thm as Thm (der, {thy_ref, tags, maxidx, shyps, hyps, tpairs, prop})) =
 | 
| 12982 | 1500  | 
(case Term.rename_abs pat obj prop of  | 
| 15531 | 1501  | 
NONE => thm  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1502  | 
| SOME prop' => Thm (der,  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
1503  | 
      {thy_ref = thy_ref,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
1504  | 
tags = tags,  | 
| 12982 | 1505  | 
maxidx = maxidx,  | 
1506  | 
hyps = hyps,  | 
|
1507  | 
shyps = shyps,  | 
|
| 
13658
 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 
berghofe 
parents: 
13642 
diff
changeset
 | 
1508  | 
tpairs = tpairs,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1509  | 
prop = prop'}));  | 
| 
10416
 
5b33e732e459
- Moved rewriting functions to meta_simplifier.ML
 
berghofe 
parents: 
10348 
diff
changeset
 | 
1510  | 
|
| 0 | 1511  | 
|
| 
44108
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1512  | 
(* strip_apply f B A strips off all assumptions/parameters from A  | 
| 0 | 1513  | 
introduced by lifting over B, and applies f to remaining part of A*)  | 
1514  | 
fun strip_apply f =  | 
|
| 
44108
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1515  | 
  let fun strip (Const ("==>", _) $ _  $ B1)
 | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1516  | 
                (Const ("==>", _) $ A2 $ B2) = Logic.mk_implies (A2, strip B1 B2)
 | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1517  | 
        | strip ((c as Const ("all", _)) $ Abs (_, _, t1))
 | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1518  | 
                (      Const ("all", _)  $ Abs (a, T, t2)) = c $ Abs (a, T, strip t1 t2)
 | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1519  | 
| strip _ A = f A  | 
| 0 | 1520  | 
in strip end;  | 
1521  | 
||
| 
44108
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1522  | 
fun strip_lifted (Const ("==>", _) $ _ $ B1)
 | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1523  | 
                 (Const ("==>", _) $ _ $ B2) = strip_lifted B1 B2
 | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1524  | 
  | strip_lifted (Const ("all", _) $ Abs (_, _, t1))
 | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1525  | 
                 (Const ("all", _) $ Abs (_, _, t2)) = strip_lifted t1 t2
 | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1526  | 
| strip_lifted _ A = A;  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1527  | 
|
| 0 | 1528  | 
(*Use the alist to rename all bound variables and some unknowns in a term  | 
1529  | 
dpairs = current disagreement pairs; tpairs = permanent ones (flexflex);  | 
|
1530  | 
Preserves unknowns in tpairs and on lhs of dpairs. *)  | 
|
| 
44108
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1531  | 
fun rename_bvs [] _ _ _ _ = K I  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1532  | 
| rename_bvs al dpairs tpairs B As =  | 
| 20330 | 1533  | 
let  | 
1534  | 
val add_var = fold_aterms (fn Var ((x, _), _) => insert (op =) x | _ => I);  | 
|
1535  | 
val vids = []  | 
|
1536  | 
|> fold (add_var o fst) dpairs  | 
|
1537  | 
|> fold (add_var o fst) tpairs  | 
|
1538  | 
|> fold (add_var o snd) tpairs;  | 
|
| 
44108
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1539  | 
val vids' = fold (add_var o strip_lifted B) As [];  | 
| 250 | 1540  | 
(*unknowns appearing elsewhere be preserved!*)  | 
| 
44108
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1541  | 
val al' = distinct ((op =) o pairself fst)  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1542  | 
(filter_out (fn (x, y) =>  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1543  | 
not (member (op =) vids' x) orelse  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1544  | 
member (op =) vids x orelse member (op =) vids y) al);  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1545  | 
val unchanged = filter_out (AList.defined (op =) al') vids';  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1546  | 
fun del_clashing clash xs _ [] qs =  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1547  | 
if clash then del_clashing false xs xs qs [] else qs  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1548  | 
| del_clashing clash xs ys ((p as (x, y)) :: ps) qs =  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1549  | 
if member (op =) ys y  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1550  | 
then del_clashing true (x :: xs) (x :: ys) ps qs  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1551  | 
else del_clashing clash xs (y :: ys) ps (p :: qs);  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1552  | 
val al'' = del_clashing false unchanged unchanged al' [];  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1553  | 
fun rename (t as Var ((x, i), T)) =  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1554  | 
(case AList.lookup (op =) al'' x of  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1555  | 
SOME y => Var ((y, i), T)  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1556  | 
| NONE => t)  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1557  | 
| rename (Abs (x, T, t)) =  | 
| 18944 | 1558  | 
Abs (the_default x (AList.lookup (op =) al x), T, rename t)  | 
| 
44108
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1559  | 
| rename (f $ t) = rename f $ rename t  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1560  | 
| rename t = t;  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1561  | 
fun strip_ren f Ai = f rename B Ai  | 
| 20330 | 1562  | 
in strip_ren end;  | 
| 0 | 1563  | 
|
1564  | 
(*Function to rename bounds/unknowns in the argument, lifted over B*)  | 
|
| 
44108
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1565  | 
fun rename_bvars dpairs =  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1566  | 
rename_bvs (List.foldr Term.match_bvars [] dpairs) dpairs;  | 
| 0 | 1567  | 
|
1568  | 
||
1569  | 
(*** RESOLUTION ***)  | 
|
1570  | 
||
| 
721
 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 
lcp 
parents: 
678 
diff
changeset
 | 
1571  | 
(** Lifting optimizations **)  | 
| 
 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 
lcp 
parents: 
678 
diff
changeset
 | 
1572  | 
|
| 0 | 1573  | 
(*strip off pairs of assumptions/parameters in parallel -- they are  | 
1574  | 
identical because of lifting*)  | 
|
| 250 | 1575  | 
fun strip_assums2 (Const("==>", _) $ _ $ B1,
 | 
1576  | 
                   Const("==>", _) $ _ $ B2) = strip_assums2 (B1,B2)
 | 
|
| 0 | 1577  | 
  | strip_assums2 (Const("all",_)$Abs(a,T,t1),
 | 
| 250 | 1578  | 
                   Const("all",_)$Abs(_,_,t2)) =
 | 
| 0 | 1579  | 
let val (B1,B2) = strip_assums2 (t1,t2)  | 
1580  | 
in (Abs(a,T,B1), Abs(a,T,B2)) end  | 
|
1581  | 
| strip_assums2 BB = BB;  | 
|
1582  | 
||
1583  | 
||
| 
721
 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 
lcp 
parents: 
678 
diff
changeset
 | 
1584  | 
(*Faster normalization: skip assumptions that were lifted over*)  | 
| 
 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 
lcp 
parents: 
678 
diff
changeset
 | 
1585  | 
fun norm_term_skip env 0 t = Envir.norm_term env t  | 
| 32032 | 1586  | 
  | norm_term_skip env n (Const ("all", _) $ Abs (a, T, t)) =
 | 
1587  | 
let  | 
|
| 32035 | 1588  | 
val T' = Envir.subst_type (Envir.type_env env) T  | 
| 32032 | 1589  | 
(*Must instantiate types of parameters because they are flattened;  | 
1590  | 
this could be a NEW parameter*)  | 
|
1591  | 
in Term.all T' $ Abs (a, T', norm_term_skip env n t) end  | 
|
1592  | 
  | norm_term_skip env n (Const ("==>", _) $ A $ B) =
 | 
|
1593  | 
Logic.mk_implies (A, norm_term_skip env (n - 1) B)  | 
|
| 32784 | 1594  | 
| norm_term_skip _ _ _ = error "norm_term_skip: too few assumptions??";  | 
| 
721
 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 
lcp 
parents: 
678 
diff
changeset
 | 
1595  | 
|
| 
 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 
lcp 
parents: 
678 
diff
changeset
 | 
1596  | 
|
| 0 | 1597  | 
(*Composition of object rule r=(A1...Am/B) with proof state s=(B1...Bn/C)  | 
| 250 | 1598  | 
Unifies B with Bi, replacing subgoal i (1 <= i <= n)  | 
| 0 | 1599  | 
If match then forbid instantiations in proof state  | 
1600  | 
If lifted then shorten the dpair using strip_assums2.  | 
|
1601  | 
If eres_flg then simultaneously proves A1 by assumption.  | 
|
| 250 | 1602  | 
nsubgoal is the number of new subgoals (written m above).  | 
| 0 | 1603  | 
Curried so that resolution calls dest_state only once.  | 
1604  | 
*)  | 
|
| 4270 | 1605  | 
local exception COMPOSE  | 
| 0 | 1606  | 
in  | 
| 18486 | 1607  | 
fun bicompose_aux flatten match (state, (stpairs, Bs, Bi, C), lifted)  | 
| 0 | 1608  | 
(eres_flg, orule, nsubgoal) =  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1609  | 
 let val Thm (sder, {maxidx=smax, shyps=sshyps, hyps=shyps, ...}) = state
 | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1610  | 
     and Thm (rder, {maxidx=rmax, shyps=rshyps, hyps=rhyps,
 | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1611  | 
tpairs=rtpairs, prop=rprop,...}) = orule  | 
| 1529 | 1612  | 
(*How many hyps to skip over during normalization*)  | 
| 21576 | 1613  | 
and nlift = Logic.count_prems (strip_all_body Bi) + (if eres_flg then ~1 else 0)  | 
| 
24143
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
1614  | 
val thy = Theory.deref (merge_thys2 state orule);  | 
| 0 | 1615  | 
(** Add new theorem with prop = '[| Bs; As |] ==> C' to thq **)  | 
| 32032 | 1616  | 
fun addth A (As, oldAs, rder', n) ((env, tpairs), thq) =  | 
| 250 | 1617  | 
let val normt = Envir.norm_term env;  | 
1618  | 
(*perform minimal copying here by examining env*)  | 
|
| 
13658
 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 
berghofe 
parents: 
13642 
diff
changeset
 | 
1619  | 
val (ntpairs, normp) =  | 
| 
 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 
berghofe 
parents: 
13642 
diff
changeset
 | 
1620  | 
if Envir.is_empty env then (tpairs, (Bs @ As, C))  | 
| 250 | 1621  | 
else  | 
1622  | 
let val ntps = map (pairself normt) tpairs  | 
|
| 19861 | 1623  | 
in if Envir.above env smax then  | 
| 1238 | 1624  | 
(*no assignments in state; normalize the rule only*)  | 
1625  | 
if lifted  | 
|
| 
13658
 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 
berghofe 
parents: 
13642 
diff
changeset
 | 
1626  | 
then (ntps, (Bs @ map (norm_term_skip env nlift) As, C))  | 
| 
 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 
berghofe 
parents: 
13642 
diff
changeset
 | 
1627  | 
else (ntps, (Bs @ map normt As, C))  | 
| 1529 | 1628  | 
else if match then raise COMPOSE  | 
| 250 | 1629  | 
else (*normalize the new rule fully*)  | 
| 
13658
 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 
berghofe 
parents: 
13642 
diff
changeset
 | 
1630  | 
(ntps, (map normt (Bs @ As), normt C))  | 
| 250 | 1631  | 
end  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1632  | 
val th =  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1633  | 
Thm (deriv_rule2  | 
| 11518 | 1634  | 
((if Envir.is_empty env then I  | 
| 19861 | 1635  | 
else if Envir.above env smax then  | 
| 37309 | 1636  | 
(fn f => fn der => f (Proofterm.norm_proof' env der))  | 
| 11518 | 1637  | 
else  | 
| 37309 | 1638  | 
curry op oo (Proofterm.norm_proof' env))  | 
1639  | 
(Proofterm.bicompose_proof flatten Bs oldAs As A n (nlift+1))) rder' sder,  | 
|
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1640  | 
                {tags = [],
 | 
| 32032 | 1641  | 
maxidx = Envir.maxidx_of env,  | 
| 
26640
 
92e6d3ec91bd
simplified handling of sorts, removed unnecessary universal witness;
 
wenzelm 
parents: 
26631 
diff
changeset
 | 
1642  | 
shyps = Envir.insert_sorts env (Sorts.union rshyps sshyps),  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1643  | 
hyps = union_hyps rhyps shyps,  | 
| 
13658
 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 
berghofe 
parents: 
13642 
diff
changeset
 | 
1644  | 
tpairs = ntpairs,  | 
| 
24143
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
1645  | 
prop = Logic.list_implies normp,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1646  | 
thy_ref = Theory.check_thy thy})  | 
| 19475 | 1647  | 
in Seq.cons th thq end handle COMPOSE => thq;  | 
| 
13658
 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 
berghofe 
parents: 
13642 
diff
changeset
 | 
1648  | 
val (rAs,B) = Logic.strip_prems(nsubgoal, [], rprop)  | 
| 0 | 1649  | 
       handle TERM _ => raise THM("bicompose: rule", 0, [orule,state]);
 | 
1650  | 
(*Modify assumptions, deleting n-th if n>0 for e-resolution*)  | 
|
1651  | 
fun newAs(As0, n, dpairs, tpairs) =  | 
|
| 11518 | 1652  | 
let val (As1, rder') =  | 
| 25939 | 1653  | 
if not lifted then (As0, rder)  | 
| 
44108
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1654  | 
else  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1655  | 
let val rename = rename_bvars dpairs tpairs B As0  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1656  | 
in (map (rename strip_apply) As0,  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1657  | 
deriv_rule1 (Proofterm.map_proof_terms (rename K) I) rder)  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
1658  | 
end;  | 
| 18486 | 1659  | 
in (map (if flatten then (Logic.flatten_params n) else I) As1, As1, rder', n)  | 
| 250 | 1660  | 
handle TERM _ =>  | 
1661  | 
          raise THM("bicompose: 1st premise", 0, [orule])
 | 
|
| 0 | 1662  | 
end;  | 
| 2147 | 1663  | 
val env = Envir.empty(Int.max(rmax,smax));  | 
| 0 | 1664  | 
val BBi = if lifted then strip_assums2(B,Bi) else (B,Bi);  | 
1665  | 
val dpairs = BBi :: (rtpairs@stpairs);  | 
|
| 
30554
 
73f8bd5f0af8
substantial speedup of assumption and elim-resolution: Logic.assum_problems refrains from eager application of parameters, discriminate via Term.could_unify before invoking full unification (assumes terms in beta-normal form, as do higher term net operations anyway);
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
1666  | 
|
| 
 
73f8bd5f0af8
substantial speedup of assumption and elim-resolution: Logic.assum_problems refrains from eager application of parameters, discriminate via Term.could_unify before invoking full unification (assumes terms in beta-normal form, as do higher term net operations anyway);
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
1667  | 
(*elim-resolution: try each assumption in turn*)  | 
| 
 
73f8bd5f0af8
substantial speedup of assumption and elim-resolution: Logic.assum_problems refrains from eager application of parameters, discriminate via Term.could_unify before invoking full unification (assumes terms in beta-normal form, as do higher term net operations anyway);
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
1668  | 
     fun eres [] = raise THM ("bicompose: no premises", 0, [orule, state])
 | 
| 
 
73f8bd5f0af8
substantial speedup of assumption and elim-resolution: Logic.assum_problems refrains from eager application of parameters, discriminate via Term.could_unify before invoking full unification (assumes terms in beta-normal form, as do higher term net operations anyway);
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
1669  | 
| eres (A1 :: As) =  | 
| 
 
73f8bd5f0af8
substantial speedup of assumption and elim-resolution: Logic.assum_problems refrains from eager application of parameters, discriminate via Term.could_unify before invoking full unification (assumes terms in beta-normal form, as do higher term net operations anyway);
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
1670  | 
let  | 
| 
 
73f8bd5f0af8
substantial speedup of assumption and elim-resolution: Logic.assum_problems refrains from eager application of parameters, discriminate via Term.could_unify before invoking full unification (assumes terms in beta-normal form, as do higher term net operations anyway);
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
1671  | 
val A = SOME A1;  | 
| 
30556
 
7be15917f3fa
eq_assumption: slightly more efficient by checking (open) result of Logic.assum_problems directly;
 
wenzelm 
parents: 
30554 
diff
changeset
 | 
1672  | 
val (close, asms, concl) = Logic.assum_problems (nlift + 1, A1);  | 
| 
 
7be15917f3fa
eq_assumption: slightly more efficient by checking (open) result of Logic.assum_problems directly;
 
wenzelm 
parents: 
30554 
diff
changeset
 | 
1673  | 
val concl' = close concl;  | 
| 
30554
 
73f8bd5f0af8
substantial speedup of assumption and elim-resolution: Logic.assum_problems refrains from eager application of parameters, discriminate via Term.could_unify before invoking full unification (assumes terms in beta-normal form, as do higher term net operations anyway);
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
1674  | 
fun tryasms [] _ = Seq.empty  | 
| 
30556
 
7be15917f3fa
eq_assumption: slightly more efficient by checking (open) result of Logic.assum_problems directly;
 
wenzelm 
parents: 
30554 
diff
changeset
 | 
1675  | 
| tryasms (asm :: rest) n =  | 
| 
 
7be15917f3fa
eq_assumption: slightly more efficient by checking (open) result of Logic.assum_problems directly;
 
wenzelm 
parents: 
30554 
diff
changeset
 | 
1676  | 
if Term.could_unify (asm, concl) then  | 
| 
 
7be15917f3fa
eq_assumption: slightly more efficient by checking (open) result of Logic.assum_problems directly;
 
wenzelm 
parents: 
30554 
diff
changeset
 | 
1677  | 
let val asm' = close asm in  | 
| 
 
7be15917f3fa
eq_assumption: slightly more efficient by checking (open) result of Logic.assum_problems directly;
 
wenzelm 
parents: 
30554 
diff
changeset
 | 
1678  | 
(case Seq.pull (Unify.unifiers (thy, env, (asm', concl') :: dpairs)) of  | 
| 
30554
 
73f8bd5f0af8
substantial speedup of assumption and elim-resolution: Logic.assum_problems refrains from eager application of parameters, discriminate via Term.could_unify before invoking full unification (assumes terms in beta-normal form, as do higher term net operations anyway);
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
1679  | 
NONE => tryasms rest (n + 1)  | 
| 
 
73f8bd5f0af8
substantial speedup of assumption and elim-resolution: Logic.assum_problems refrains from eager application of parameters, discriminate via Term.could_unify before invoking full unification (assumes terms in beta-normal form, as do higher term net operations anyway);
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
1680  | 
| cell as SOME ((_, tpairs), _) =>  | 
| 
30556
 
7be15917f3fa
eq_assumption: slightly more efficient by checking (open) result of Logic.assum_problems directly;
 
wenzelm 
parents: 
30554 
diff
changeset
 | 
1681  | 
Seq.it_right (addth A (newAs (As, n, [BBi, (concl', asm')], tpairs)))  | 
| 
30554
 
73f8bd5f0af8
substantial speedup of assumption and elim-resolution: Logic.assum_problems refrains from eager application of parameters, discriminate via Term.could_unify before invoking full unification (assumes terms in beta-normal form, as do higher term net operations anyway);
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
1682  | 
(Seq.make (fn () => cell),  | 
| 
 
73f8bd5f0af8
substantial speedup of assumption and elim-resolution: Logic.assum_problems refrains from eager application of parameters, discriminate via Term.could_unify before invoking full unification (assumes terms in beta-normal form, as do higher term net operations anyway);
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
1683  | 
Seq.make (fn () => Seq.pull (tryasms rest (n + 1)))))  | 
| 
 
73f8bd5f0af8
substantial speedup of assumption and elim-resolution: Logic.assum_problems refrains from eager application of parameters, discriminate via Term.could_unify before invoking full unification (assumes terms in beta-normal form, as do higher term net operations anyway);
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
1684  | 
end  | 
| 
 
73f8bd5f0af8
substantial speedup of assumption and elim-resolution: Logic.assum_problems refrains from eager application of parameters, discriminate via Term.could_unify before invoking full unification (assumes terms in beta-normal form, as do higher term net operations anyway);
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
1685  | 
else tryasms rest (n + 1);  | 
| 
30556
 
7be15917f3fa
eq_assumption: slightly more efficient by checking (open) result of Logic.assum_problems directly;
 
wenzelm 
parents: 
30554 
diff
changeset
 | 
1686  | 
in tryasms asms 1 end;  | 
| 
30554
 
73f8bd5f0af8
substantial speedup of assumption and elim-resolution: Logic.assum_problems refrains from eager application of parameters, discriminate via Term.could_unify before invoking full unification (assumes terms in beta-normal form, as do higher term net operations anyway);
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
1687  | 
|
| 0 | 1688  | 
(*ordinary resolution*)  | 
| 
30554
 
73f8bd5f0af8
substantial speedup of assumption and elim-resolution: Logic.assum_problems refrains from eager application of parameters, discriminate via Term.could_unify before invoking full unification (assumes terms in beta-normal form, as do higher term net operations anyway);
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
1689  | 
fun res () =  | 
| 
 
73f8bd5f0af8
substantial speedup of assumption and elim-resolution: Logic.assum_problems refrains from eager application of parameters, discriminate via Term.could_unify before invoking full unification (assumes terms in beta-normal form, as do higher term net operations anyway);
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
1690  | 
(case Seq.pull (Unify.unifiers (thy, env, dpairs)) of  | 
| 
 
73f8bd5f0af8
substantial speedup of assumption and elim-resolution: Logic.assum_problems refrains from eager application of parameters, discriminate via Term.could_unify before invoking full unification (assumes terms in beta-normal form, as do higher term net operations anyway);
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
1691  | 
NONE => Seq.empty  | 
| 
 
73f8bd5f0af8
substantial speedup of assumption and elim-resolution: Logic.assum_problems refrains from eager application of parameters, discriminate via Term.could_unify before invoking full unification (assumes terms in beta-normal form, as do higher term net operations anyway);
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
1692  | 
| cell as SOME ((_, tpairs), _) =>  | 
| 
 
73f8bd5f0af8
substantial speedup of assumption and elim-resolution: Logic.assum_problems refrains from eager application of parameters, discriminate via Term.could_unify before invoking full unification (assumes terms in beta-normal form, as do higher term net operations anyway);
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
1693  | 
Seq.it_right (addth NONE (newAs (rev rAs, 0, [BBi], tpairs)))  | 
| 
 
73f8bd5f0af8
substantial speedup of assumption and elim-resolution: Logic.assum_problems refrains from eager application of parameters, discriminate via Term.could_unify before invoking full unification (assumes terms in beta-normal form, as do higher term net operations anyway);
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
1694  | 
(Seq.make (fn () => cell), Seq.empty));  | 
| 
 
73f8bd5f0af8
substantial speedup of assumption and elim-resolution: Logic.assum_problems refrains from eager application of parameters, discriminate via Term.could_unify before invoking full unification (assumes terms in beta-normal form, as do higher term net operations anyway);
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
1695  | 
in  | 
| 
 
73f8bd5f0af8
substantial speedup of assumption and elim-resolution: Logic.assum_problems refrains from eager application of parameters, discriminate via Term.could_unify before invoking full unification (assumes terms in beta-normal form, as do higher term net operations anyway);
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
1696  | 
if eres_flg then eres (rev rAs) else res ()  | 
| 0 | 1697  | 
end;  | 
| 7528 | 1698  | 
end;  | 
| 0 | 1699  | 
|
| 
18501
 
915105af2e80
turned bicompose_no_flatten into compose_no_flatten, without elimination;
 
wenzelm 
parents: 
18486 
diff
changeset
 | 
1700  | 
fun compose_no_flatten match (orule, nsubgoal) i state =  | 
| 
 
915105af2e80
turned bicompose_no_flatten into compose_no_flatten, without elimination;
 
wenzelm 
parents: 
18486 
diff
changeset
 | 
1701  | 
bicompose_aux false match (state, dest_state (state, i), false) (false, orule, nsubgoal);  | 
| 0 | 1702  | 
|
| 
18501
 
915105af2e80
turned bicompose_no_flatten into compose_no_flatten, without elimination;
 
wenzelm 
parents: 
18486 
diff
changeset
 | 
1703  | 
fun bicompose match arg i state =  | 
| 
 
915105af2e80
turned bicompose_no_flatten into compose_no_flatten, without elimination;
 
wenzelm 
parents: 
18486 
diff
changeset
 | 
1704  | 
bicompose_aux true match (state, dest_state (state,i), false) arg;  | 
| 0 | 1705  | 
|
1706  | 
(*Quick test whether rule is resolvable with the subgoal with hyps Hs  | 
|
1707  | 
and conclusion B. If eres_flg then checks 1st premise of rule also*)  | 
|
1708  | 
fun could_bires (Hs, B, eres_flg, rule) =  | 
|
| 
29269
 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 
wenzelm 
parents: 
29003 
diff
changeset
 | 
1709  | 
let fun could_reshyp (A1::_) = exists (fn H => Term.could_unify (A1, H)) Hs  | 
| 250 | 1710  | 
| could_reshyp [] = false; (*no premise -- illegal*)  | 
| 
29269
 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 
wenzelm 
parents: 
29003 
diff
changeset
 | 
1711  | 
in Term.could_unify(concl_of rule, B) andalso  | 
| 250 | 1712  | 
(not eres_flg orelse could_reshyp (prems_of rule))  | 
| 0 | 1713  | 
end;  | 
1714  | 
||
1715  | 
(*Bi-resolution of a state with a list of (flag,rule) pairs.  | 
|
1716  | 
Puts the rule above: rule/state. Renames vars in the rules. *)  | 
|
| 250 | 1717  | 
fun biresolution match brules i state =  | 
| 18035 | 1718  | 
let val (stpairs, Bs, Bi, C) = dest_state(state,i);  | 
| 18145 | 1719  | 
val lift = lift_rule (cprem_of state i);  | 
| 250 | 1720  | 
val B = Logic.strip_assums_concl Bi;  | 
1721  | 
val Hs = Logic.strip_assums_hyp Bi;  | 
|
| 22573 | 1722  | 
val compose = bicompose_aux true match (state, (stpairs, Bs, Bi, C), true);  | 
| 4270 | 1723  | 
fun res [] = Seq.empty  | 
| 250 | 1724  | 
| res ((eres_flg, rule)::brules) =  | 
| 
13642
 
a3d97348ceb6
added failure trace information to pattern unification
 
nipkow 
parents: 
13629 
diff
changeset
 | 
1725  | 
if !Pattern.trace_unify_fail orelse  | 
| 
 
a3d97348ceb6
added failure trace information to pattern unification
 
nipkow 
parents: 
13629 
diff
changeset
 | 
1726  | 
could_bires (Hs, B, eres_flg, rule)  | 
| 4270 | 1727  | 
then Seq.make (*delay processing remainder till needed*)  | 
| 22573 | 1728  | 
(fn()=> SOME(compose (eres_flg, lift rule, nprems_of rule),  | 
| 250 | 1729  | 
res brules))  | 
1730  | 
else res brules  | 
|
| 4270 | 1731  | 
in Seq.flat (res brules) end;  | 
| 0 | 1732  | 
|
1733  | 
||
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1734  | 
|
| 2509 | 1735  | 
(*** Oracles ***)  | 
1736  | 
||
| 28290 | 1737  | 
(* oracle rule *)  | 
1738  | 
||
1739  | 
fun invoke_oracle thy_ref1 name oracle arg =  | 
|
| 28624 | 1740  | 
  let val Cterm {thy_ref = thy_ref2, t = prop, T, maxidx, sorts} = oracle arg in
 | 
| 28290 | 1741  | 
if T <> propT then  | 
1742  | 
      raise THM ("Oracle's result must have type prop: " ^ name, 0, [])
 | 
|
1743  | 
else  | 
|
| 37309 | 1744  | 
let val (ora, prf) = Proofterm.oracle_proof name prop in  | 
| 
32059
 
7991cdf10a54
support for arbitrarity nested future proofs -- replaced crude order by explicit normalization (which might loop for bad dependencies);
 
wenzelm 
parents: 
32035 
diff
changeset
 | 
1745  | 
Thm (make_deriv [] [ora] [] prf,  | 
| 
28804
 
5d3b1df16353
refined notion of derivation, consiting of promises and proof_body;
 
wenzelm 
parents: 
28675 
diff
changeset
 | 
1746  | 
         {thy_ref = Theory.merge_refs (thy_ref1, thy_ref2),
 | 
| 
 
5d3b1df16353
refined notion of derivation, consiting of promises and proof_body;
 
wenzelm 
parents: 
28675 
diff
changeset
 | 
1747  | 
tags = [],  | 
| 
 
5d3b1df16353
refined notion of derivation, consiting of promises and proof_body;
 
wenzelm 
parents: 
28675 
diff
changeset
 | 
1748  | 
maxidx = maxidx,  | 
| 
 
5d3b1df16353
refined notion of derivation, consiting of promises and proof_body;
 
wenzelm 
parents: 
28675 
diff
changeset
 | 
1749  | 
shyps = sorts,  | 
| 
 
5d3b1df16353
refined notion of derivation, consiting of promises and proof_body;
 
wenzelm 
parents: 
28675 
diff
changeset
 | 
1750  | 
hyps = [],  | 
| 
 
5d3b1df16353
refined notion of derivation, consiting of promises and proof_body;
 
wenzelm 
parents: 
28675 
diff
changeset
 | 
1751  | 
tpairs = [],  | 
| 
 
5d3b1df16353
refined notion of derivation, consiting of promises and proof_body;
 
wenzelm 
parents: 
28675 
diff
changeset
 | 
1752  | 
prop = prop})  | 
| 
 
5d3b1df16353
refined notion of derivation, consiting of promises and proof_body;
 
wenzelm 
parents: 
28675 
diff
changeset
 | 
1753  | 
end  | 
| 3812 | 1754  | 
end;  | 
1755  | 
||
| 
32590
 
95f4f08f950f
replaced opaque signature matching by plain old abstype (again, cf. ac4498f95d1c) -- this recovers pretty printing in SML/NJ and Poly/ML 5.3;
 
wenzelm 
parents: 
32198 
diff
changeset
 | 
1756  | 
end;  | 
| 
 
95f4f08f950f
replaced opaque signature matching by plain old abstype (again, cf. ac4498f95d1c) -- this recovers pretty printing in SML/NJ and Poly/ML 5.3;
 
wenzelm 
parents: 
32198 
diff
changeset
 | 
1757  | 
end;  | 
| 
 
95f4f08f950f
replaced opaque signature matching by plain old abstype (again, cf. ac4498f95d1c) -- this recovers pretty printing in SML/NJ and Poly/ML 5.3;
 
wenzelm 
parents: 
32198 
diff
changeset
 | 
1758  | 
end;  | 
| 
 
95f4f08f950f
replaced opaque signature matching by plain old abstype (again, cf. ac4498f95d1c) -- this recovers pretty printing in SML/NJ and Poly/ML 5.3;
 
wenzelm 
parents: 
32198 
diff
changeset
 | 
1759  | 
|
| 28290 | 1760  | 
|
1761  | 
(* authentic derivation names *)  | 
|
1762  | 
||
| 33522 | 1763  | 
structure Oracles = Theory_Data  | 
| 28290 | 1764  | 
(  | 
| 
33095
 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 
wenzelm 
parents: 
33092 
diff
changeset
 | 
1765  | 
type T = unit Name_Space.table;  | 
| 33159 | 1766  | 
val empty : T = Name_Space.empty_table "oracle";  | 
| 28290 | 1767  | 
val extend = I;  | 
| 33522 | 1768  | 
fun merge data : T = Name_Space.merge_tables data;  | 
| 28290 | 1769  | 
);  | 
1770  | 
||
| 
42358
 
b47d41d9f4b5
Name_Space: proper configuration options long_names, short_names, unique_names instead of former unsynchronized references;
 
wenzelm 
parents: 
41700 
diff
changeset
 | 
1771  | 
fun extern_oracles ctxt =  | 
| 42360 | 1772  | 
map #1 (Name_Space.extern_table ctxt (Oracles.get (Proof_Context.theory_of ctxt)));  | 
| 28290 | 1773  | 
|
| 
30288
 
a32700e45ab3
Thm.add_oracle interface: replaced old bstring by binding;
 
wenzelm 
parents: 
29636 
diff
changeset
 | 
1774  | 
fun add_oracle (b, oracle) thy =  | 
| 28290 | 1775  | 
let  | 
1776  | 
val naming = Sign.naming_of thy;  | 
|
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
1777  | 
val (name, tab') =  | 
| 
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
1778  | 
Name_Space.define (Proof_Context.init_global thy) true naming (b, ()) (Oracles.get thy);  | 
| 
30288
 
a32700e45ab3
Thm.add_oracle interface: replaced old bstring by binding;
 
wenzelm 
parents: 
29636 
diff
changeset
 | 
1779  | 
val thy' = Oracles.put tab' thy;  | 
| 28290 | 1780  | 
in ((name, invoke_oracle (Theory.check_thy thy') name oracle), thy') end;  | 
1781  | 
||
| 0 | 1782  | 
end;  | 
| 1503 | 1783  | 
|
| 32104 | 1784  | 
structure Basic_Thm: BASIC_THM = Thm;  | 
1785  | 
open Basic_Thm;  |