| author | desharna | 
| Mon, 23 May 2022 10:23:33 +0200 | |
| changeset 75460 | 7c2fe41f5ee8 | 
| parent 74645 | 30eba7f9a8e9 | 
| child 77730 | 4a174bea55e2 | 
| 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,  | 
| 59589 | 6  | 
derivations, theorems, inference rules (including lifting and  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
7  | 
resolution), oracles.  | 
| 0 | 8  | 
*)  | 
9  | 
||
| 
70456
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
10  | 
infix 0 RS RSN;  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
11  | 
|
| 6089 | 12  | 
signature BASIC_THM =  | 
| 59589 | 13  | 
sig  | 
| 59582 | 14  | 
type ctyp  | 
15  | 
type cterm  | 
|
16  | 
exception CTERM of string * cterm list  | 
|
17  | 
type thm  | 
|
18  | 
type conv = cterm -> thm  | 
|
19  | 
exception THM of string * int * thm list  | 
|
| 
70456
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
20  | 
val RSN: thm * (int * thm) -> thm  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
21  | 
val RS: thm * thm -> thm  | 
| 59582 | 22  | 
end;  | 
23  | 
||
24  | 
signature THM =  | 
|
25  | 
sig  | 
|
26  | 
include BASIC_THM  | 
|
| 1160 | 27  | 
(*certified types*)  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
28  | 
val typ_of: ctyp -> typ  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59591 
diff
changeset
 | 
29  | 
val global_ctyp_of: theory -> typ -> ctyp  | 
| 
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59591 
diff
changeset
 | 
30  | 
val ctyp_of: Proof.context -> typ -> ctyp  | 
| 59582 | 31  | 
val dest_ctyp: ctyp -> ctyp list  | 
| 70159 | 32  | 
val dest_ctypN: int -> ctyp -> ctyp  | 
33  | 
val dest_ctyp0: ctyp -> ctyp  | 
|
34  | 
val dest_ctyp1: ctyp -> ctyp  | 
|
| 70156 | 35  | 
val make_ctyp: ctyp -> ctyp list -> ctyp  | 
| 1160 | 36  | 
(*certified terms*)  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
37  | 
val term_of: cterm -> term  | 
| 59586 | 38  | 
val typ_of_cterm: cterm -> typ  | 
39  | 
val ctyp_of_cterm: cterm -> ctyp  | 
|
40  | 
val maxidx_of_cterm: cterm -> int  | 
|
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59591 
diff
changeset
 | 
41  | 
val global_cterm_of: theory -> term -> cterm  | 
| 
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59591 
diff
changeset
 | 
42  | 
val cterm_of: Proof.context -> term -> cterm  | 
| 
59969
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59917 
diff
changeset
 | 
43  | 
val renamed_term: term -> cterm -> cterm  | 
| 74270 | 44  | 
val fast_term_ord: cterm ord  | 
45  | 
val term_ord: cterm ord  | 
|
| 59582 | 46  | 
val dest_comb: cterm -> cterm * cterm  | 
47  | 
val dest_fun: cterm -> cterm  | 
|
48  | 
val dest_arg: cterm -> cterm  | 
|
49  | 
val dest_fun2: cterm -> cterm  | 
|
50  | 
val dest_arg1: cterm -> cterm  | 
|
| 
74525
 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 
wenzelm 
parents: 
74518 
diff
changeset
 | 
51  | 
val dest_abs_fresh: string -> cterm -> cterm * cterm  | 
| 
 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 
wenzelm 
parents: 
74518 
diff
changeset
 | 
52  | 
val dest_abs_global: cterm -> cterm * cterm  | 
| 
60952
 
762cb38a3147
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60951 
diff
changeset
 | 
53  | 
val rename_tvar: indexname -> ctyp -> ctyp  | 
| 
60951
 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60949 
diff
changeset
 | 
54  | 
val var: indexname * ctyp -> cterm  | 
| 59582 | 55  | 
val apply: cterm -> cterm -> cterm  | 
56  | 
val lambda_name: string * cterm -> cterm -> cterm  | 
|
57  | 
val lambda: cterm -> cterm -> cterm  | 
|
58  | 
val adjust_maxidx_cterm: int -> cterm -> cterm  | 
|
59  | 
val incr_indexes_cterm: int -> cterm -> cterm  | 
|
| 74282 | 60  | 
val match: cterm * cterm -> ctyp TVars.table * cterm Vars.table  | 
61  | 
val first_order_match: cterm * cterm -> ctyp TVars.table * cterm Vars.table  | 
|
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
62  | 
(*theorems*)  | 
| 
74241
 
eb265f54e3ce
more efficient operations: traverse hyps only when required;
 
wenzelm 
parents: 
74235 
diff
changeset
 | 
63  | 
  val fold_terms: {hyps: bool} -> (term -> 'a -> 'a) -> thm -> 'a -> 'a
 | 
| 74244 | 64  | 
  val fold_atomic_ctyps: {hyps: bool} -> (typ -> bool) -> (ctyp -> 'a -> 'a) -> thm -> 'a -> 'a
 | 
65  | 
  val fold_atomic_cterms: {hyps: bool} -> (term -> bool) -> (cterm -> 'a -> 'a) -> thm -> 'a -> 'a
 | 
|
| 59582 | 66  | 
val terms_of_tpairs: (term * term) list -> term list  | 
67  | 
val full_prop_of: thm -> term  | 
|
| 65458 | 68  | 
val theory_id: thm -> Context.theory_id  | 
69  | 
val theory_name: thm -> string  | 
|
| 59582 | 70  | 
val maxidx_of: thm -> int  | 
71  | 
val maxidx_thm: thm -> int -> int  | 
|
| 61039 | 72  | 
val shyps_of: thm -> sort Ord_List.T  | 
| 59582 | 73  | 
val hyps_of: thm -> term list  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
74  | 
val prop_of: thm -> term  | 
| 59582 | 75  | 
val tpairs_of: thm -> (term * term) list  | 
| 16656 | 76  | 
val concl_of: thm -> term  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
77  | 
val prems_of: thm -> term list  | 
| 
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
78  | 
val nprems_of: thm -> int  | 
| 59582 | 79  | 
val no_prems: thm -> bool  | 
80  | 
val major_prem_of: thm -> term  | 
|
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
81  | 
val cprop_of: thm -> cterm  | 
| 18145 | 82  | 
val cprem_of: thm -> int -> cterm  | 
| 
69101
 
991a3feaf270
more direct implementation of distinct_subgoals_tac -- potentially more efficient;
 
wenzelm 
parents: 
68692 
diff
changeset
 | 
83  | 
val cconcl_of: thm -> cterm  | 
| 
 
991a3feaf270
more direct implementation of distinct_subgoals_tac -- potentially more efficient;
 
wenzelm 
parents: 
68692 
diff
changeset
 | 
84  | 
val cprems_of: thm -> cterm list  | 
| 60949 | 85  | 
val chyps_of: thm -> cterm list  | 
| 74270 | 86  | 
val thm_ord: thm ord  | 
| 61051 | 87  | 
exception CONTEXT of string * ctyp list * cterm list * thm list * Context.generic option  | 
88  | 
val theory_of_cterm: cterm -> theory  | 
|
89  | 
val theory_of_thm: thm -> theory  | 
|
| 67697 | 90  | 
val trim_context_ctyp: ctyp -> ctyp  | 
| 61051 | 91  | 
val trim_context_cterm: cterm -> cterm  | 
| 70156 | 92  | 
val transfer_ctyp: theory -> ctyp -> ctyp  | 
| 61051 | 93  | 
val transfer_cterm: theory -> cterm -> cterm  | 
| 38709 | 94  | 
val transfer: theory -> thm -> thm  | 
| 67649 | 95  | 
val transfer': Proof.context -> thm -> thm  | 
96  | 
val transfer'': Context.generic -> thm -> thm  | 
|
| 
70310
 
c82f59c47aaf
clarified transfer_morphism: implicit join_certificate, e.g. relevant for complex cascades of morphisms such as class locale interpretation;
 
wenzelm 
parents: 
70159 
diff
changeset
 | 
97  | 
val join_transfer: theory -> thm -> thm  | 
| 70313 | 98  | 
val join_transfer_context: Proof.context * thm -> Proof.context * thm  | 
| 
59969
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59917 
diff
changeset
 | 
99  | 
val renamed_prop: term -> thm -> thm  | 
| 38709 | 100  | 
val weaken: cterm -> thm -> thm  | 
101  | 
val weaken_sorts: sort list -> cterm -> cterm  | 
|
| 
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
 | 
102  | 
val proof_bodies_of: thm list -> proof_body list  | 
| 32725 | 103  | 
val proof_body_of: thm -> proof_body  | 
104  | 
val proof_of: thm -> proof  | 
|
| 71088 | 105  | 
val reconstruct_proof_of: thm -> Proofterm.proof  | 
| 66168 | 106  | 
val consolidate: thm list -> unit  | 
| 71011 | 107  | 
val expose_proofs: theory -> thm list -> unit  | 
| 
71017
 
c85efa2be619
expose derivations more thoroughly, notably for locale/class reasoning;
 
wenzelm 
parents: 
71012 
diff
changeset
 | 
108  | 
val expose_proof: theory -> thm -> unit  | 
| 32725 | 109  | 
val future: thm future -> cterm -> thm  | 
| 
70595
 
2ae7e33c950f
clarified thm_id vs. thm_node/thm: retain theory_name;
 
wenzelm 
parents: 
70593 
diff
changeset
 | 
110  | 
val thm_deps: thm -> Proofterm.thm Ord_List.T  | 
| 71527 | 111  | 
val extra_shyps: thm -> sort list  | 
112  | 
val strip_shyps: thm -> thm  | 
|
| 
64568
 
a504a3dec35a
more careful derivation_closed / close_derivation;
 
wenzelm 
parents: 
63858 
diff
changeset
 | 
113  | 
val derivation_closed: thm -> bool  | 
| 
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
 | 
114  | 
val derivation_name: thm -> string  | 
| 70554 | 115  | 
val derivation_id: thm -> Proofterm.thm_id option  | 
| 
70543
 
33749040b6f8
clarified derivation_name vs. raw_derivation_name;
 
wenzelm 
parents: 
70517 
diff
changeset
 | 
116  | 
val raw_derivation_name: thm -> string  | 
| 
70915
 
bd4d37edfee4
clarified expand_proof/expand_name: allow more detailed control via thm_header;
 
wenzelm 
parents: 
70908 
diff
changeset
 | 
117  | 
val expand_name: thm -> Proofterm.thm_header -> string option  | 
| 70494 | 118  | 
val name_derivation: string * Position.T -> thm -> thm  | 
119  | 
val close_derivation: Position.T -> thm -> thm  | 
|
| 
70480
 
1a1b7d7f24bb
explicit check of left-over constraints from different theory, e.g. due to lack of Thm.trim_context;
 
wenzelm 
parents: 
70475 
diff
changeset
 | 
120  | 
val trim_context: thm -> thm  | 
| 
28675
 
fb68c0767004
renamed get_axiom_i to axiom, removed obsolete get_axiom;
 
wenzelm 
parents: 
28648 
diff
changeset
 | 
121  | 
val axiom: theory -> string -> thm  | 
| 
70923
 
98d9b78b7f47
clarified axiom_table: uniform space (e.g. like consts), e.g. relevant for export of HOL-ex.Join_Theory;
 
wenzelm 
parents: 
70915 
diff
changeset
 | 
122  | 
val all_axioms_of: theory -> (string * thm) list  | 
| 28017 | 123  | 
val get_tags: thm -> Properties.T  | 
124  | 
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
 | 
125  | 
val norm_proof: thm -> thm  | 
| 20261 | 126  | 
val adjust_maxidx_thm: int -> thm -> thm  | 
| 70458 | 127  | 
(*type classes*)  | 
128  | 
val the_classrel: theory -> class * class -> thm  | 
|
129  | 
val the_arity: theory -> string * sort list * class -> thm  | 
|
130  | 
val classrel_proof: theory -> class * class -> proof  | 
|
131  | 
val arity_proof: theory -> string * sort list * class -> proof  | 
|
| 70454 | 132  | 
(*oracles*)  | 
133  | 
  val add_oracle: binding * ('a -> cterm) -> theory -> (string * ('a -> thm)) * theory
 | 
|
| 70560 | 134  | 
val oracle_space: theory -> Name_Space.T  | 
135  | 
val pretty_oracle: Proof.context -> string -> Pretty.T  | 
|
| 70454 | 136  | 
val extern_oracles: bool -> Proof.context -> (Markup.T * xstring) list  | 
| 70565 | 137  | 
val check_oracle: Proof.context -> xstring * Position.T -> string  | 
| 59589 | 138  | 
(*inference rules*)  | 
| 38709 | 139  | 
val assume: cterm -> thm  | 
140  | 
val implies_intr: cterm -> thm -> thm  | 
|
141  | 
val implies_elim: thm -> thm -> thm  | 
|
142  | 
val forall_intr: cterm -> thm -> thm  | 
|
143  | 
val forall_elim: cterm -> thm -> thm  | 
|
144  | 
val reflexive: cterm -> thm  | 
|
145  | 
val symmetric: thm -> thm  | 
|
146  | 
val transitive: thm -> thm -> thm  | 
|
147  | 
val beta_conversion: bool -> conv  | 
|
148  | 
val eta_conversion: conv  | 
|
149  | 
val eta_long_conversion: conv  | 
|
150  | 
val abstract_rule: string -> cterm -> thm -> thm  | 
|
151  | 
val combination: thm -> thm -> thm  | 
|
152  | 
val equal_intr: thm -> thm -> thm  | 
|
153  | 
val equal_elim: thm -> thm -> thm  | 
|
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
154  | 
val solve_constraints: thm -> thm  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
58946 
diff
changeset
 | 
155  | 
val flexflex_rule: Proof.context option -> thm -> thm Seq.seq  | 
| 74266 | 156  | 
val generalize: Names.set * Names.set -> int -> thm -> thm  | 
157  | 
val generalize_cterm: Names.set * Names.set -> int -> cterm -> cterm  | 
|
158  | 
val generalize_ctyp: Names.set -> int -> ctyp -> ctyp  | 
|
| 74282 | 159  | 
val instantiate: ctyp TVars.table * cterm Vars.table -> thm -> thm  | 
| 
74577
 
d4829a7333e2
clarified instantiation: local beta reduction after substitution, as for Envir.expand_term_defs;
 
wenzelm 
parents: 
74561 
diff
changeset
 | 
160  | 
val instantiate_beta: ctyp TVars.table * cterm Vars.table -> thm -> thm  | 
| 74282 | 161  | 
val instantiate_cterm: ctyp TVars.table * cterm Vars.table -> cterm -> cterm  | 
| 
74577
 
d4829a7333e2
clarified instantiation: local beta reduction after substitution, as for Envir.expand_term_defs;
 
wenzelm 
parents: 
74561 
diff
changeset
 | 
162  | 
val instantiate_beta_cterm: ctyp TVars.table * cterm Vars.table -> cterm -> cterm  | 
| 38709 | 163  | 
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
 | 
164  | 
val of_class: ctyp * class -> thm  | 
| 
36768
 
46be86127972
just one version of Thm.unconstrainT, which affects all variables;
 
wenzelm 
parents: 
36744 
diff
changeset
 | 
165  | 
val unconstrainT: thm -> thm  | 
| 74266 | 166  | 
val varifyT_global': TFrees.set -> thm -> ((string * sort) * indexname) list * thm  | 
| 38709 | 167  | 
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
 | 
168  | 
val legacy_freezeT: thm -> thm  | 
| 
70456
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
169  | 
val plain_prop_of: thm -> term  | 
| 
59969
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59917 
diff
changeset
 | 
170  | 
val dest_state: thm * int -> (term * term) list * term list * term * term  | 
| 38709 | 171  | 
val lift_rule: cterm -> thm -> thm  | 
172  | 
val incr_indexes: int -> thm -> thm  | 
|
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
58946 
diff
changeset
 | 
173  | 
val assumption: Proof.context option -> int -> thm -> thm Seq.seq  | 
| 31945 | 174  | 
val eq_assumption: int -> thm -> thm  | 
175  | 
val rotate_rule: int -> int -> thm -> thm  | 
|
176  | 
val permute_prems: int -> int -> thm -> thm  | 
|
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
58946 
diff
changeset
 | 
177  | 
  val bicompose: Proof.context option -> {flatten: bool, match: bool, incremented: bool} ->
 | 
| 
52223
 
5bb6ae8acb87
tuned signature -- more explicit flags for low-level Thm.bicompose;
 
wenzelm 
parents: 
52222 
diff
changeset
 | 
178  | 
bool * thm * int -> int -> thm -> thm Seq.seq  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
58946 
diff
changeset
 | 
179  | 
val biresolution: Proof.context option -> bool -> (bool * thm) list -> int -> thm -> thm Seq.seq  | 
| 
70475
 
98b6da301e13
backed out changeset 1b8858f4c393: odd problems e.g. in CAVA_LTL_Modelchecker;
 
wenzelm 
parents: 
70472 
diff
changeset
 | 
180  | 
val thynames_of_arity: theory -> string * class -> string list  | 
| 
70456
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
181  | 
val add_classrel: thm -> theory -> theory  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
182  | 
val add_arity: thm -> theory -> theory  | 
| 6089 | 183  | 
end;  | 
184  | 
||
| 
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
 | 
185  | 
structure Thm: THM =  | 
| 0 | 186  | 
struct  | 
| 250 | 187  | 
|
| 
387
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
188  | 
(*** Certified terms and types ***)  | 
| 
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
189  | 
|
| 250 | 190  | 
(** certified types **)  | 
191  | 
||
| 
70453
 
492cb3aaa562
simplified module structure: back to plain datatype (see 95f4f08f950f and 70019ab5e57f);
 
wenzelm 
parents: 
70404 
diff
changeset
 | 
192  | 
datatype ctyp = Ctyp of {cert: Context.certificate, T: typ, maxidx: int, sorts: sort Ord_List.T};
 | 
| 250 | 193  | 
|
194  | 
fun typ_of (Ctyp {T, ...}) = T;
 | 
|
195  | 
||
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59591 
diff
changeset
 | 
196  | 
fun global_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
 | 
197  | 
let  | 
| 
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
198  | 
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
 | 
199  | 
val maxidx = Term.maxidx_of_typ T;  | 
| 
26640
 
92e6d3ec91bd
simplified handling of sorts, removed unnecessary universal witness;
 
wenzelm 
parents: 
26631 
diff
changeset
 | 
200  | 
val sorts = Sorts.insert_typ T [];  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
201  | 
  in Ctyp {cert = Context.Certificate thy, T = T, maxidx = maxidx, sorts = sorts} end;
 | 
| 250 | 202  | 
|
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59591 
diff
changeset
 | 
203  | 
val ctyp_of = global_ctyp_of o Proof_Context.theory_of;  | 
| 
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59591 
diff
changeset
 | 
204  | 
|
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
205  | 
fun dest_ctyp (Ctyp {cert, T = Type (_, Ts), maxidx, sorts}) =
 | 
| 
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
206  | 
      map (fn T => Ctyp {cert = cert, T = T, maxidx = maxidx, sorts = sorts}) Ts
 | 
| 16679 | 207  | 
  | dest_ctyp cT = raise TYPE ("dest_ctyp", [typ_of cT], []);
 | 
| 15087 | 208  | 
|
| 70159 | 209  | 
fun dest_ctypN n (Ctyp {cert, T, maxidx, sorts}) =
 | 
210  | 
  let fun err () = raise TYPE ("dest_ctypN", [T], []) in
 | 
|
| 
70153
 
8a23083ac011
prefer exception TYPE, e.g. when used within conversion;
 
wenzelm 
parents: 
70150 
diff
changeset
 | 
211  | 
(case T of  | 
| 
 
8a23083ac011
prefer exception TYPE, e.g. when used within conversion;
 
wenzelm 
parents: 
70150 
diff
changeset
 | 
212  | 
Type (_, Ts) =>  | 
| 70159 | 213  | 
        Ctyp {cert = cert, T = nth Ts n handle General.Subscript => err (),
 | 
| 
70153
 
8a23083ac011
prefer exception TYPE, e.g. when used within conversion;
 
wenzelm 
parents: 
70150 
diff
changeset
 | 
214  | 
maxidx = maxidx, sorts = sorts}  | 
| 
 
8a23083ac011
prefer exception TYPE, e.g. when used within conversion;
 
wenzelm 
parents: 
70150 
diff
changeset
 | 
215  | 
| _ => err ())  | 
| 
 
8a23083ac011
prefer exception TYPE, e.g. when used within conversion;
 
wenzelm 
parents: 
70150 
diff
changeset
 | 
216  | 
end;  | 
| 70150 | 217  | 
|
| 70159 | 218  | 
val dest_ctyp0 = dest_ctypN 0;  | 
219  | 
val dest_ctyp1 = dest_ctypN 1;  | 
|
220  | 
||
| 70156 | 221  | 
fun join_certificate_ctyp (Ctyp {cert, ...}) cert0 = Context.join_certificate (cert0, cert);
 | 
222  | 
fun union_sorts_ctyp (Ctyp {sorts, ...}) sorts0 = Sorts.union sorts0 sorts;
 | 
|
223  | 
fun maxidx_ctyp (Ctyp {maxidx, ...}) maxidx0 = Int.max (maxidx0, maxidx);
 | 
|
224  | 
||
225  | 
fun make_ctyp (Ctyp {cert, T, maxidx = _, sorts = _}) cargs =
 | 
|
226  | 
let  | 
|
227  | 
val As = map typ_of cargs;  | 
|
228  | 
    fun err () = raise TYPE ("make_ctyp", T :: As, []);
 | 
|
229  | 
in  | 
|
230  | 
(case T of  | 
|
231  | 
Type (a, args) =>  | 
|
232  | 
        Ctyp {
 | 
|
233  | 
cert = fold join_certificate_ctyp cargs cert,  | 
|
234  | 
maxidx = fold maxidx_ctyp cargs ~1,  | 
|
235  | 
sorts = fold union_sorts_ctyp cargs [],  | 
|
236  | 
T = if length args = length cargs then Type (a, As) else err ()}  | 
|
237  | 
| _ => err ())  | 
|
238  | 
end;  | 
|
239  | 
||
| 
229
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
225 
diff
changeset
 | 
240  | 
|
| 
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
225 
diff
changeset
 | 
241  | 
|
| 250 | 242  | 
(** certified terms **)  | 
| 
229
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
225 
diff
changeset
 | 
243  | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
244  | 
(*certified terms with checked typ, maxidx, and sorts*)  | 
| 
70453
 
492cb3aaa562
simplified module structure: back to plain datatype (see 95f4f08f950f and 70019ab5e57f);
 
wenzelm 
parents: 
70404 
diff
changeset
 | 
245  | 
datatype cterm =  | 
| 
 
492cb3aaa562
simplified module structure: back to plain datatype (see 95f4f08f950f and 70019ab5e57f);
 
wenzelm 
parents: 
70404 
diff
changeset
 | 
246  | 
  Cterm of {cert: Context.certificate, t: term, T: typ, maxidx: int, sorts: sort Ord_List.T};
 | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
16352 
diff
changeset
 | 
247  | 
|
| 22584 | 248  | 
exception CTERM of string * cterm list;  | 
| 16679 | 249  | 
|
| 250 | 250  | 
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
 | 
251  | 
|
| 59586 | 252  | 
fun typ_of_cterm (Cterm {T, ...}) = T;
 | 
253  | 
||
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
254  | 
fun ctyp_of_cterm (Cterm {cert, T, maxidx, sorts, ...}) =
 | 
| 
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
255  | 
  Ctyp {cert = cert, T = T, maxidx = maxidx, sorts = sorts};
 | 
| 2671 | 256  | 
|
| 59586 | 257  | 
fun maxidx_of_cterm (Cterm {maxidx, ...}) = maxidx;
 | 
258  | 
||
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59591 
diff
changeset
 | 
259  | 
fun global_cterm_of thy tm =  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
260  | 
let  | 
| 18969 | 261  | 
val (t, T, maxidx) = Sign.certify_term thy tm;  | 
| 
26640
 
92e6d3ec91bd
simplified handling of sorts, removed unnecessary universal witness;
 
wenzelm 
parents: 
26631 
diff
changeset
 | 
262  | 
val sorts = Sorts.insert_term t [];  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
263  | 
  in Cterm {cert = Context.Certificate 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
 | 
264  | 
|
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59591 
diff
changeset
 | 
265  | 
val cterm_of = global_cterm_of o Proof_Context.theory_of;  | 
| 
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59591 
diff
changeset
 | 
266  | 
|
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
267  | 
fun join_certificate0 (Cterm {cert = cert1, ...}, Cterm {cert = cert2, ...}) =
 | 
| 
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
268  | 
Context.join_certificate (cert1, cert2);  | 
| 60315 | 269  | 
|
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
270  | 
fun renamed_term t' (Cterm {cert, t, T, maxidx, sorts}) =
 | 
| 
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
271  | 
  if t aconv t' then Cterm {cert = cert, t = t', T = T, maxidx = maxidx, sorts = sorts}
 | 
| 
59969
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59917 
diff
changeset
 | 
272  | 
  else raise TERM ("renamed_term: terms disagree", [t, t']);
 | 
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59917 
diff
changeset
 | 
273  | 
|
| 74270 | 274  | 
val fast_term_ord = Term_Ord.fast_term_ord o apply2 term_of;  | 
275  | 
val term_ord = Term_Ord.term_ord o apply2 term_of;  | 
|
276  | 
||
| 
20580
 
6fb75df09253
added dest_arg, i.e. a tuned version of #2 o dest_comb;
 
wenzelm 
parents: 
20548 
diff
changeset
 | 
277  | 
|
| 22909 | 278  | 
(* destructors *)  | 
279  | 
||
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
280  | 
fun dest_comb (Cterm {t = c $ a, T, cert, maxidx, sorts}) =
 | 
| 22909 | 281  | 
let val A = Term.argument_type_of c 0 in  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
282  | 
        (Cterm {t = c, T = A --> T, cert = cert, maxidx = maxidx, sorts = sorts},
 | 
| 
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
283  | 
         Cterm {t = a, T = A, cert = cert, maxidx = maxidx, sorts = sorts})
 | 
| 
1493
 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 
clasohm 
parents: 
1460 
diff
changeset
 | 
284  | 
end  | 
| 22584 | 285  | 
  | 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
 | 
286  | 
|
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
287  | 
fun dest_fun (Cterm {t = c $ _, T, cert, maxidx, sorts}) =
 | 
| 22909 | 288  | 
let val A = Term.argument_type_of c 0  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
289  | 
      in Cterm {t = c, T = A --> T, cert = cert, maxidx = maxidx, sorts = sorts} end
 | 
| 22909 | 290  | 
  | dest_fun ct = raise CTERM ("dest_fun", [ct]);
 | 
291  | 
||
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
292  | 
fun dest_arg (Cterm {t = c $ a, T = _, cert, maxidx, sorts}) =
 | 
| 22909 | 293  | 
let val A = Term.argument_type_of c 0  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
294  | 
      in Cterm {t = a, T = A, cert = cert, maxidx = maxidx, sorts = sorts} end
 | 
| 22584 | 295  | 
  | 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
 | 
296  | 
|
| 22909 | 297  | 
|
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
298  | 
fun dest_fun2 (Cterm {t = c $ _ $ _, T, cert, maxidx, sorts}) =
 | 
| 22909 | 299  | 
let  | 
300  | 
val A = Term.argument_type_of c 0;  | 
|
301  | 
val B = Term.argument_type_of c 1;  | 
|
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
302  | 
      in Cterm {t = c, T = A --> B --> T, cert = cert, maxidx = maxidx, sorts = sorts} end
 | 
| 22909 | 303  | 
  | dest_fun2 ct = raise CTERM ("dest_fun2", [ct]);
 | 
304  | 
||
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
305  | 
fun dest_arg1 (Cterm {t = c $ a $ _, T = _, cert, maxidx, sorts}) =
 | 
| 22909 | 306  | 
let val A = Term.argument_type_of c 0  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
307  | 
      in Cterm {t = a, T = A, cert = cert, maxidx = maxidx, sorts = sorts} end
 | 
| 22909 | 308  | 
  | dest_arg1 ct = raise CTERM ("dest_arg1", [ct]);
 | 
| 20673 | 309  | 
|
| 
74525
 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 
wenzelm 
parents: 
74518 
diff
changeset
 | 
310  | 
fun gen_dest_abs dest ct =  | 
| 
 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 
wenzelm 
parents: 
74518 
diff
changeset
 | 
311  | 
(case ct of  | 
| 
 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 
wenzelm 
parents: 
74518 
diff
changeset
 | 
312  | 
    Cterm {t = t as Abs _, T = Type ("fun", [_, U]), cert, maxidx, sorts} =>
 | 
| 
 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 
wenzelm 
parents: 
74518 
diff
changeset
 | 
313  | 
let  | 
| 
 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 
wenzelm 
parents: 
74518 
diff
changeset
 | 
314  | 
val ((x', T), t') = dest t;  | 
| 
 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 
wenzelm 
parents: 
74518 
diff
changeset
 | 
315  | 
        val v = Cterm {t = Free (x', T), T = T, cert = cert, maxidx = maxidx, sorts = sorts};
 | 
| 
 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 
wenzelm 
parents: 
74518 
diff
changeset
 | 
316  | 
        val body = Cterm {t = t', T = U, cert = cert, maxidx = maxidx, sorts = sorts};
 | 
| 
 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 
wenzelm 
parents: 
74518 
diff
changeset
 | 
317  | 
in (v, body) end  | 
| 
 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 
wenzelm 
parents: 
74518 
diff
changeset
 | 
318  | 
  | _ => raise CTERM ("dest_abs", [ct]));
 | 
| 74518 | 319  | 
|
| 
74525
 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 
wenzelm 
parents: 
74518 
diff
changeset
 | 
320  | 
val dest_abs_fresh = gen_dest_abs o Term.dest_abs_fresh;  | 
| 
 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 
wenzelm 
parents: 
74518 
diff
changeset
 | 
321  | 
val dest_abs_global = gen_dest_abs Term.dest_abs_global;  | 
| 
1493
 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 
clasohm 
parents: 
1460 
diff
changeset
 | 
322  | 
|
| 22909 | 323  | 
|
324  | 
(* constructors *)  | 
|
325  | 
||
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
326  | 
fun rename_tvar (a, i) (Ctyp {cert, T, maxidx, sorts}) =
 | 
| 
60952
 
762cb38a3147
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60951 
diff
changeset
 | 
327  | 
let  | 
| 
 
762cb38a3147
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60951 
diff
changeset
 | 
328  | 
val S =  | 
| 
 
762cb38a3147
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60951 
diff
changeset
 | 
329  | 
(case T of  | 
| 
 
762cb38a3147
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60951 
diff
changeset
 | 
330  | 
TFree (_, S) => S  | 
| 
 
762cb38a3147
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60951 
diff
changeset
 | 
331  | 
| TVar (_, S) => S  | 
| 
 
762cb38a3147
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60951 
diff
changeset
 | 
332  | 
      | _ => raise TYPE ("rename_tvar: no variable", [T], []));
 | 
| 
 
762cb38a3147
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60951 
diff
changeset
 | 
333  | 
    val _ = if i < 0 then raise TYPE ("rename_tvar: bad index", [TVar ((a, i), S)], []) else ();
 | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
334  | 
  in Ctyp {cert = cert, T = TVar ((a, i), S), maxidx = Int.max (i, maxidx), sorts = sorts} end;
 | 
| 
60952
 
762cb38a3147
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60951 
diff
changeset
 | 
335  | 
|
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
336  | 
fun var ((x, i), Ctyp {cert, T, maxidx, sorts}) =
 | 
| 
60951
 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60949 
diff
changeset
 | 
337  | 
  if i < 0 then raise TERM ("var: bad index", [Var ((x, i), T)])
 | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
338  | 
  else Cterm {cert = cert, t = Var ((x, i), T), T = T, maxidx = Int.max (i, maxidx), sorts = sorts};
 | 
| 
60951
 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60949 
diff
changeset
 | 
339  | 
|
| 
46497
 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 
wenzelm 
parents: 
46475 
diff
changeset
 | 
340  | 
fun apply  | 
| 16656 | 341  | 
  (cf as Cterm {t = f, T = Type ("fun", [dty, rty]), maxidx = maxidx1, sorts = sorts1, ...})
 | 
342  | 
  (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
 | 
343  | 
if T = dty then  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
344  | 
      Cterm {cert = join_certificate0 (cf, cx),
 | 
| 16656 | 345  | 
t = f $ x,  | 
346  | 
T = rty,  | 
|
347  | 
maxidx = Int.max (maxidx1, maxidx2),  | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
348  | 
sorts = Sorts.union sorts1 sorts2}  | 
| 
46497
 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 
wenzelm 
parents: 
46475 
diff
changeset
 | 
349  | 
      else raise CTERM ("apply: types don't agree", [cf, cx])
 | 
| 
 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 
wenzelm 
parents: 
46475 
diff
changeset
 | 
350  | 
  | apply cf cx = raise CTERM ("apply: first arg is not a function", [cf, cx]);
 | 
| 250 | 351  | 
|
| 
46497
 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 
wenzelm 
parents: 
46475 
diff
changeset
 | 
352  | 
fun lambda_name  | 
| 32198 | 353  | 
  (x, ct1 as Cterm {t = t1, T = T1, maxidx = maxidx1, sorts = sorts1, ...})
 | 
| 16656 | 354  | 
  (ct2 as Cterm {t = t2, T = T2, maxidx = maxidx2, sorts = sorts2, ...}) =
 | 
| 32198 | 355  | 
let val t = Term.lambda_name (x, t1) t2 in  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
356  | 
      Cterm {cert = join_certificate0 (ct1, ct2),
 | 
| 16656 | 357  | 
t = t, T = T1 --> T2,  | 
358  | 
maxidx = Int.max (maxidx1, maxidx2),  | 
|
359  | 
sorts = Sorts.union sorts1 sorts2}  | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
360  | 
end;  | 
| 
229
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
225 
diff
changeset
 | 
361  | 
|
| 
46497
 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 
wenzelm 
parents: 
46475 
diff
changeset
 | 
362  | 
fun lambda t u = lambda_name ("", t) u;
 | 
| 32198 | 363  | 
|
| 
20580
 
6fb75df09253
added dest_arg, i.e. a tuned version of #2 o dest_comb;
 
wenzelm 
parents: 
20548 
diff
changeset
 | 
364  | 
|
| 22909 | 365  | 
(* indexes *)  | 
366  | 
||
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
367  | 
fun adjust_maxidx_cterm i (ct as Cterm {cert, t, T, maxidx, sorts}) =
 | 
| 
20580
 
6fb75df09253
added dest_arg, i.e. a tuned version of #2 o dest_comb;
 
wenzelm 
parents: 
20548 
diff
changeset
 | 
368  | 
if maxidx = i then ct  | 
| 
 
6fb75df09253
added dest_arg, i.e. a tuned version of #2 o dest_comb;
 
wenzelm 
parents: 
20548 
diff
changeset
 | 
369  | 
else if maxidx < i then  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
370  | 
    Cterm {maxidx = i, cert = cert, t = t, T = T, sorts = sorts}
 | 
| 
20580
 
6fb75df09253
added dest_arg, i.e. a tuned version of #2 o dest_comb;
 
wenzelm 
parents: 
20548 
diff
changeset
 | 
371  | 
else  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
372  | 
    Cterm {maxidx = Int.max (maxidx_of_term t, i), cert = cert, t = t, T = T, sorts = sorts};
 | 
| 
20580
 
6fb75df09253
added dest_arg, i.e. a tuned version of #2 o dest_comb;
 
wenzelm 
parents: 
20548 
diff
changeset
 | 
373  | 
|
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
374  | 
fun incr_indexes_cterm i (ct as Cterm {cert, t, T, maxidx, sorts}) =
 | 
| 22909 | 375  | 
  if i < 0 then raise CTERM ("negative increment", [ct])
 | 
376  | 
else if i = 0 then ct  | 
|
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
377  | 
  else Cterm {cert = cert, t = Logic.incr_indexes ([], [], i) t,
 | 
| 22909 | 378  | 
T = Logic.incr_tvar i T, maxidx = maxidx + i, sorts = sorts};  | 
379  | 
||
380  | 
||
| 2509 | 381  | 
|
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
382  | 
(*** Derivations and Theorems ***)  | 
| 
229
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
225 
diff
changeset
 | 
383  | 
|
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
384  | 
(* sort constraints *)  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
385  | 
|
| 
70480
 
1a1b7d7f24bb
explicit check of left-over constraints from different theory, e.g. due to lack of Thm.trim_context;
 
wenzelm 
parents: 
70475 
diff
changeset
 | 
386  | 
type constraint = {theory: theory, typ: typ, sort: sort};
 | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
387  | 
|
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
388  | 
local  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
389  | 
|
| 70586 | 390  | 
val constraint_ord : constraint ord =  | 
| 
70480
 
1a1b7d7f24bb
explicit check of left-over constraints from different theory, e.g. due to lack of Thm.trim_context;
 
wenzelm 
parents: 
70475 
diff
changeset
 | 
391  | 
Context.theory_id_ord o apply2 (Context.theory_id o #theory)  | 
| 73860 | 392  | 
||| Term_Ord.typ_ord o apply2 #typ  | 
393  | 
||| Term_Ord.sort_ord o apply2 #sort;  | 
|
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
394  | 
|
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
395  | 
val smash_atyps =  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
396  | 
map_atyps (fn TVar (_, S) => Term.aT S | TFree (_, S) => Term.aT S | T => T);  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
397  | 
|
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
398  | 
in  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
399  | 
|
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
400  | 
val union_constraints = Ord_List.union constraint_ord;  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
401  | 
|
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
402  | 
fun insert_constraints thy (T, S) =  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
403  | 
let  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
404  | 
val ignored =  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
405  | 
S = [] orelse  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
406  | 
(case T of  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
407  | 
TFree (_, S') => S = S'  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
408  | 
| TVar (_, S') => S = S'  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
409  | 
| _ => false);  | 
| 
70480
 
1a1b7d7f24bb
explicit check of left-over constraints from different theory, e.g. due to lack of Thm.trim_context;
 
wenzelm 
parents: 
70475 
diff
changeset
 | 
410  | 
in  | 
| 
 
1a1b7d7f24bb
explicit check of left-over constraints from different theory, e.g. due to lack of Thm.trim_context;
 
wenzelm 
parents: 
70475 
diff
changeset
 | 
411  | 
if ignored then I  | 
| 
 
1a1b7d7f24bb
explicit check of left-over constraints from different theory, e.g. due to lack of Thm.trim_context;
 
wenzelm 
parents: 
70475 
diff
changeset
 | 
412  | 
    else Ord_List.insert constraint_ord {theory = thy, typ = smash_atyps T, sort = S}
 | 
| 
 
1a1b7d7f24bb
explicit check of left-over constraints from different theory, e.g. due to lack of Thm.trim_context;
 
wenzelm 
parents: 
70475 
diff
changeset
 | 
413  | 
end;  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
414  | 
|
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
415  | 
fun insert_constraints_env thy env =  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
416  | 
let  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
417  | 
val tyenv = Envir.type_env env;  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
418  | 
fun insert ([], _) = I  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
419  | 
| insert (S, T) = insert_constraints thy (Envir.norm_type tyenv T, S);  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
420  | 
in tyenv |> Vartab.fold (insert o #2) end;  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
421  | 
|
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
422  | 
end;  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
423  | 
|
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
424  | 
|
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
425  | 
(* datatype thm *)  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
426  | 
|
| 
70453
 
492cb3aaa562
simplified module structure: back to plain datatype (see 95f4f08f950f and 70019ab5e57f);
 
wenzelm 
parents: 
70404 
diff
changeset
 | 
427  | 
datatype thm = Thm of  | 
| 40124 | 428  | 
deriv * (*derivation*)  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
429  | 
 {cert: Context.certificate,    (*background theory certificate*)
 | 
| 40124 | 430  | 
tags: Properties.T, (*additional annotations/comments*)  | 
431  | 
maxidx: int, (*maximum index of any Var or TVar*)  | 
|
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
432  | 
constraints: constraint Ord_List.T, (*implicit proof obligations for sort constraints*)  | 
| 40124 | 433  | 
shyps: sort Ord_List.T, (*sort hypotheses*)  | 
434  | 
hyps: term Ord_List.T, (*hypotheses*)  | 
|
435  | 
tpairs: (term * term) list, (*flex-flex pairs*)  | 
|
436  | 
prop: term} (*conclusion*)  | 
|
| 28624 | 437  | 
and deriv = Deriv of  | 
| 39687 | 438  | 
 {promises: (serial * thm future) Ord_List.T,
 | 
| 
70453
 
492cb3aaa562
simplified module structure: back to plain datatype (see 95f4f08f950f and 70019ab5e57f);
 
wenzelm 
parents: 
70404 
diff
changeset
 | 
439  | 
body: Proofterm.proof_body};  | 
| 0 | 440  | 
|
| 23601 | 441  | 
type conv = cterm -> thm;  | 
442  | 
||
| 16725 | 443  | 
(*errors involving theorems*)  | 
444  | 
exception THM of string * int * thm list;  | 
|
| 
13658
 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 
berghofe 
parents: 
13642 
diff
changeset
 | 
445  | 
|
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
446  | 
fun rep_thm (Thm (_, args)) = args;  | 
| 0 | 447  | 
|
| 
74241
 
eb265f54e3ce
more efficient operations: traverse hyps only when required;
 
wenzelm 
parents: 
74235 
diff
changeset
 | 
448  | 
fun fold_terms h f (Thm (_, {tpairs, prop, hyps, ...})) =
 | 
| 
 
eb265f54e3ce
more efficient operations: traverse hyps only when required;
 
wenzelm 
parents: 
74235 
diff
changeset
 | 
449  | 
fold (fn (t, u) => f t #> f u) tpairs #> f prop #> #hyps h ? fold f hyps;  | 
| 
31947
 
7daee3bed3af
clarified strip_shyps: proper type witnesses for present sorts;
 
wenzelm 
parents: 
31945 
diff
changeset
 | 
450  | 
|
| 74244 | 451  | 
fun fold_atomic_ctyps h g f (th as Thm (_, {cert, maxidx, shyps, ...})) =
 | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
452  | 
  let fun ctyp T = Ctyp {cert = cert, T = T, maxidx = maxidx, sorts = shyps}
 | 
| 74244 | 453  | 
in (fold_terms h o fold_types o fold_atyps) (fn T => if g T then f (ctyp T) else I) th end;  | 
| 
60952
 
762cb38a3147
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60951 
diff
changeset
 | 
454  | 
|
| 74244 | 455  | 
fun fold_atomic_cterms h g f (th as Thm (_, {cert, maxidx, shyps, ...})) =
 | 
456  | 
let  | 
|
457  | 
    fun cterm t T = Cterm {cert = cert, t = t, T = T, maxidx = maxidx, sorts = shyps};
 | 
|
458  | 
fun apply t T = if g t then f (cterm t T) else I;  | 
|
459  | 
in  | 
|
| 
74241
 
eb265f54e3ce
more efficient operations: traverse hyps only when required;
 
wenzelm 
parents: 
74235 
diff
changeset
 | 
460  | 
(fold_terms h o fold_aterms)  | 
| 74244 | 461  | 
(fn t as Const (_, T) => apply t T  | 
462  | 
| t as Free (_, T) => apply t T  | 
|
463  | 
| t as Var (_, T) => apply t T  | 
|
| 60818 | 464  | 
| _ => I) th  | 
465  | 
end;  | 
|
466  | 
||
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
467  | 
|
| 16725 | 468  | 
fun terms_of_tpairs tpairs = fold_rev (fn (t, u) => cons t o cons u) tpairs [];  | 
469  | 
||
470  | 
fun eq_tpairs ((t, u), (t', u')) = t aconv t' andalso u aconv u';  | 
|
| 18944 | 471  | 
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
 | 
472  | 
val maxidx_tpairs = fold (fn (t, u) => Term.maxidx_term t #> Term.maxidx_term u);  | 
| 16725 | 473  | 
|
474  | 
fun attach_tpairs tpairs prop =  | 
|
475  | 
Logic.list_implies (map Logic.mk_equals tpairs, prop);  | 
|
476  | 
||
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
477  | 
fun full_prop_of (Thm (_, {tpairs, prop, ...})) = attach_tpairs tpairs prop;
 | 
| 16945 | 478  | 
|
| 
70456
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
479  | 
|
| 39687 | 480  | 
val union_hyps = Ord_List.union Term_Ord.fast_term_ord;  | 
481  | 
val insert_hyps = Ord_List.insert Term_Ord.fast_term_ord;  | 
|
482  | 
val remove_hyps = Ord_List.remove Term_Ord.fast_term_ord;  | 
|
| 22365 | 483  | 
|
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
484  | 
fun join_certificate1 (Cterm {cert = cert1, ...}, Thm (_, {cert = cert2, ...})) =
 | 
| 
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
485  | 
Context.join_certificate (cert1, cert2);  | 
| 16945 | 486  | 
|
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
487  | 
fun join_certificate2 (Thm (_, {cert = cert1, ...}), Thm (_, {cert = cert2, ...})) =
 | 
| 
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
488  | 
Context.join_certificate (cert1, cert2);  | 
| 16945 | 489  | 
|
| 0 | 490  | 
|
| 22365 | 491  | 
(* basic components *)  | 
| 16135 | 492  | 
|
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
493  | 
val cert_of = #cert o rep_thm;  | 
| 65458 | 494  | 
val theory_id = Context.certificate_theory_id o cert_of;  | 
495  | 
val theory_name = Context.theory_id_name o theory_id;  | 
|
| 61050 | 496  | 
|
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
497  | 
val maxidx_of = #maxidx o rep_thm;  | 
| 19910 | 498  | 
fun maxidx_thm th i = Int.max (maxidx_of th, i);  | 
| 61039 | 499  | 
val shyps_of = #shyps o rep_thm;  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
500  | 
val hyps_of = #hyps o rep_thm;  | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
501  | 
val prop_of = #prop o rep_thm;  | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
502  | 
val tpairs_of = #tpairs o rep_thm;  | 
| 0 | 503  | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
504  | 
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
 | 
505  | 
val prems_of = Logic.strip_imp_prems o prop_of;  | 
| 21576 | 506  | 
val nprems_of = Logic.count_prems o prop_of;  | 
| 74509 | 507  | 
val no_prems = Logic.no_prems o prop_of;  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
508  | 
|
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
509  | 
fun major_prem_of th =  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
510  | 
(case prems_of th of  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
511  | 
prem :: _ => Logic.strip_assums_concl prem  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
512  | 
  | [] => 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
 | 
513  | 
|
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
514  | 
fun cprop_of (Thm (_, {cert, maxidx, shyps, prop, ...})) =
 | 
| 
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
515  | 
  Cterm {cert = cert, maxidx = maxidx, T = propT, t = prop, sorts = shyps};
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
516  | 
|
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
517  | 
fun cprem_of (th as Thm (_, {cert, maxidx, shyps, prop, ...})) i =
 | 
| 
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
518  | 
  Cterm {cert = cert, maxidx = maxidx, T = propT, sorts = shyps,
 | 
| 18145 | 519  | 
    t = Logic.nth_prem (i, prop) handle TERM _ => raise THM ("cprem_of", i, [th])};
 | 
| 18035 | 520  | 
|
| 
69101
 
991a3feaf270
more direct implementation of distinct_subgoals_tac -- potentially more efficient;
 
wenzelm 
parents: 
68692 
diff
changeset
 | 
521  | 
fun cconcl_of (th as Thm (_, {cert, maxidx, shyps, ...})) =
 | 
| 
 
991a3feaf270
more direct implementation of distinct_subgoals_tac -- potentially more efficient;
 
wenzelm 
parents: 
68692 
diff
changeset
 | 
522  | 
  Cterm {cert = cert, maxidx = maxidx, T = propT, sorts = shyps, t = concl_of th};
 | 
| 
 
991a3feaf270
more direct implementation of distinct_subgoals_tac -- potentially more efficient;
 
wenzelm 
parents: 
68692 
diff
changeset
 | 
523  | 
|
| 
 
991a3feaf270
more direct implementation of distinct_subgoals_tac -- potentially more efficient;
 
wenzelm 
parents: 
68692 
diff
changeset
 | 
524  | 
fun cprems_of (th as Thm (_, {cert, maxidx, shyps, ...})) =
 | 
| 
 
991a3feaf270
more direct implementation of distinct_subgoals_tac -- potentially more efficient;
 
wenzelm 
parents: 
68692 
diff
changeset
 | 
525  | 
  map (fn t => Cterm {cert = cert, maxidx = maxidx, T = propT, sorts = shyps, t = t})
 | 
| 
 
991a3feaf270
more direct implementation of distinct_subgoals_tac -- potentially more efficient;
 
wenzelm 
parents: 
68692 
diff
changeset
 | 
526  | 
(prems_of th);  | 
| 
 
991a3feaf270
more direct implementation of distinct_subgoals_tac -- potentially more efficient;
 
wenzelm 
parents: 
68692 
diff
changeset
 | 
527  | 
|
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
528  | 
fun chyps_of (Thm (_, {cert, shyps, hyps, ...})) =
 | 
| 
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
529  | 
  map (fn t => Cterm {cert = cert, maxidx = ~1, T = propT, sorts = shyps, t = t}) hyps;
 | 
| 60949 | 530  | 
|
| 61051 | 531  | 
|
| 74270 | 532  | 
(* thm order: ignores theory context! *)  | 
533  | 
||
534  | 
val thm_ord =  | 
|
| 74645 | 535  | 
pointer_eq_ord  | 
536  | 
(Term_Ord.fast_term_ord o apply2 prop_of  | 
|
537  | 
||| list_ord (prod_ord Term_Ord.fast_term_ord Term_Ord.fast_term_ord) o apply2 tpairs_of  | 
|
538  | 
||| list_ord Term_Ord.fast_term_ord o apply2 hyps_of  | 
|
539  | 
||| list_ord Term_Ord.sort_ord o apply2 shyps_of);  | 
|
| 74270 | 540  | 
|
541  | 
||
| 61051 | 542  | 
(* implicit theory context *)  | 
543  | 
||
544  | 
exception CONTEXT of string * ctyp list * cterm list * thm list * Context.generic option;  | 
|
545  | 
||
546  | 
fun theory_of_cterm (ct as Cterm {cert, ...}) =
 | 
|
547  | 
Context.certificate_theory cert  | 
|
548  | 
handle ERROR msg => raise CONTEXT (msg, [], [ct], [], NONE);  | 
|
549  | 
||
550  | 
fun theory_of_thm th =  | 
|
551  | 
Context.certificate_theory (cert_of th)  | 
|
552  | 
handle ERROR msg => raise CONTEXT (msg, [], [], [th], NONE);  | 
|
553  | 
||
| 67697 | 554  | 
fun trim_context_ctyp cT =  | 
555  | 
(case cT of  | 
|
556  | 
    Ctyp {cert = Context.Certificate_Id _, ...} => cT
 | 
|
557  | 
  | Ctyp {cert = Context.Certificate thy, T, maxidx, sorts} =>
 | 
|
558  | 
      Ctyp {cert = Context.Certificate_Id (Context.theory_id thy),
 | 
|
559  | 
T = T, maxidx = maxidx, sorts = sorts});  | 
|
560  | 
||
| 61051 | 561  | 
fun trim_context_cterm ct =  | 
562  | 
(case ct of  | 
|
563  | 
    Cterm {cert = Context.Certificate_Id _, ...} => ct
 | 
|
564  | 
  | Cterm {cert = Context.Certificate thy, t, T, maxidx, sorts} =>
 | 
|
565  | 
      Cterm {cert = Context.Certificate_Id (Context.theory_id thy),
 | 
|
566  | 
t = t, T = T, maxidx = maxidx, sorts = sorts});  | 
|
567  | 
||
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
568  | 
fun trim_context_thm th =  | 
| 61048 | 569  | 
(case th of  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
570  | 
    Thm (_, {constraints = _ :: _, ...}) =>
 | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
571  | 
      raise THM ("trim_context: pending sort constraints", 0, [th])
 | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
572  | 
  | Thm (_, {cert = Context.Certificate_Id _, ...}) => th
 | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
573  | 
| Thm (der,  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
574  | 
      {cert = Context.Certificate thy, tags, maxidx, constraints = [], shyps, hyps,
 | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
575  | 
tpairs, prop}) =>  | 
| 61048 | 576  | 
Thm (der,  | 
577  | 
       {cert = Context.Certificate_Id (Context.theory_id thy),
 | 
|
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
578  | 
tags = tags, maxidx = maxidx, constraints = [], shyps = shyps, hyps = hyps,  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
579  | 
tpairs = tpairs, prop = prop}));  | 
| 61048 | 580  | 
|
| 70156 | 581  | 
fun transfer_ctyp thy' cT =  | 
582  | 
let  | 
|
583  | 
    val Ctyp {cert, T, maxidx, sorts} = cT;
 | 
|
584  | 
val _ =  | 
|
585  | 
Context.subthy_id (Context.certificate_theory_id cert, Context.theory_id thy') orelse  | 
|
586  | 
        raise CONTEXT ("Cannot transfer: not a super theory", [cT], [], [],
 | 
|
587  | 
SOME (Context.Theory thy'));  | 
|
588  | 
val cert' = Context.join_certificate (Context.Certificate thy', cert);  | 
|
589  | 
in  | 
|
590  | 
if Context.eq_certificate (cert, cert') then cT  | 
|
591  | 
    else Ctyp {cert = cert', T = T, maxidx = maxidx, sorts = sorts}
 | 
|
592  | 
end;  | 
|
593  | 
||
| 61051 | 594  | 
fun transfer_cterm thy' ct =  | 
| 3895 | 595  | 
let  | 
| 61051 | 596  | 
    val Cterm {cert, t, T, maxidx, sorts} = ct;
 | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
597  | 
val _ =  | 
| 
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
598  | 
Context.subthy_id (Context.certificate_theory_id cert, Context.theory_id thy') orelse  | 
| 61051 | 599  | 
        raise CONTEXT ("Cannot transfer: not a super theory", [], [ct], [],
 | 
600  | 
SOME (Context.Theory thy'));  | 
|
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
601  | 
val cert' = Context.join_certificate (Context.Certificate thy', cert);  | 
| 3895 | 602  | 
in  | 
| 61051 | 603  | 
if Context.eq_certificate (cert, cert') then ct  | 
604  | 
    else Cterm {cert = cert', t = t, T = T, maxidx = maxidx, sorts = sorts}
 | 
|
605  | 
end;  | 
|
606  | 
||
607  | 
fun transfer thy' th =  | 
|
608  | 
let  | 
|
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
609  | 
    val Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop}) = th;
 | 
| 61051 | 610  | 
val _ =  | 
611  | 
Context.subthy_id (Context.certificate_theory_id cert, Context.theory_id thy') orelse  | 
|
612  | 
        raise CONTEXT ("Cannot transfer: not a super theory", [], [], [th],
 | 
|
613  | 
SOME (Context.Theory thy'));  | 
|
614  | 
val cert' = Context.join_certificate (Context.Certificate thy', cert);  | 
|
615  | 
in  | 
|
616  | 
if Context.eq_certificate (cert, cert') then th  | 
|
| 16945 | 617  | 
else  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
618  | 
Thm (der,  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
619  | 
       {cert = cert',
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
620  | 
tags = tags,  | 
| 16945 | 621  | 
maxidx = maxidx,  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
622  | 
constraints = constraints,  | 
| 16945 | 623  | 
shyps = shyps,  | 
624  | 
hyps = hyps,  | 
|
625  | 
tpairs = tpairs,  | 
|
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
626  | 
prop = prop})  | 
| 3895 | 627  | 
end;  | 
| 
387
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
628  | 
|
| 67649 | 629  | 
val transfer' = transfer o Proof_Context.theory_of;  | 
630  | 
val transfer'' = transfer o Context.theory_of;  | 
|
631  | 
||
| 
70310
 
c82f59c47aaf
clarified transfer_morphism: implicit join_certificate, e.g. relevant for complex cascades of morphisms such as class locale interpretation;
 
wenzelm 
parents: 
70159 
diff
changeset
 | 
632  | 
fun join_transfer thy th =  | 
| 
71553
 
cf2406e654cf
more robust: proper transfer if Context.eq_thy_id;
 
wenzelm 
parents: 
71530 
diff
changeset
 | 
633  | 
(Context.subthy_id (theory_id th, Context.theory_id thy) ? transfer thy) th;  | 
| 
70310
 
c82f59c47aaf
clarified transfer_morphism: implicit join_certificate, e.g. relevant for complex cascades of morphisms such as class locale interpretation;
 
wenzelm 
parents: 
70159 
diff
changeset
 | 
634  | 
|
| 70313 | 635  | 
fun join_transfer_context (ctxt, th) =  | 
| 71177 | 636  | 
if Context.subthy_id (theory_id th, Context.theory_id (Proof_Context.theory_of ctxt))  | 
637  | 
then (ctxt, transfer' ctxt th)  | 
|
638  | 
else (Context.raw_transfer (theory_of_thm th) ctxt, th);  | 
|
| 70313 | 639  | 
|
| 61051 | 640  | 
|
641  | 
(* matching *)  | 
|
642  | 
||
643  | 
local  | 
|
644  | 
||
645  | 
fun gen_match match  | 
|
646  | 
    (ct1 as Cterm {t = t1, sorts = sorts1, ...},
 | 
|
647  | 
     ct2 as Cterm {t = t2, sorts = sorts2, maxidx = maxidx2, ...}) =
 | 
|
648  | 
let  | 
|
649  | 
val cert = join_certificate0 (ct1, ct2);  | 
|
650  | 
val thy = Context.certificate_theory cert  | 
|
651  | 
handle ERROR msg => raise CONTEXT (msg, [], [ct1, ct2], [], NONE);  | 
|
652  | 
val (Tinsts, tinsts) = match thy (t1, t2) (Vartab.empty, Vartab.empty);  | 
|
653  | 
val sorts = Sorts.union sorts1 sorts2;  | 
|
654  | 
fun mk_cTinst ((a, i), (S, T)) =  | 
|
655  | 
      (((a, i), S), Ctyp {T = T, cert = cert, maxidx = maxidx2, sorts = sorts});
 | 
|
656  | 
fun mk_ctinst ((x, i), (U, t)) =  | 
|
657  | 
let val T = Envir.subst_type Tinsts U in  | 
|
658  | 
        (((x, i), T), Cterm {t = t, T = T, cert = cert, maxidx = maxidx2, sorts = sorts})
 | 
|
659  | 
end;  | 
|
| 74282 | 660  | 
in  | 
661  | 
(TVars.build (Vartab.fold (TVars.add o mk_cTinst) Tinsts),  | 
|
662  | 
Vars.build (Vartab.fold (Vars.add o mk_ctinst) tinsts))  | 
|
663  | 
end;  | 
|
| 61051 | 664  | 
|
665  | 
in  | 
|
666  | 
||
667  | 
val match = gen_match Pattern.match;  | 
|
668  | 
val first_order_match = gen_match Pattern.first_order_match;  | 
|
669  | 
||
670  | 
end;  | 
|
671  | 
||
672  | 
||
| 
59969
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59917 
diff
changeset
 | 
673  | 
(*implicit alpha-conversion*)  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
674  | 
fun renamed_prop prop' (Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop})) =
 | 
| 
59969
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59917 
diff
changeset
 | 
675  | 
if prop aconv prop' then  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
676  | 
    Thm (der, {cert = cert, tags = tags, maxidx = maxidx, constraints = constraints, shyps = shyps,
 | 
| 
59969
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59917 
diff
changeset
 | 
677  | 
hyps = hyps, tpairs = tpairs, prop = prop'})  | 
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59917 
diff
changeset
 | 
678  | 
  else raise TERM ("renamed_prop: props disagree", [prop, prop']);
 | 
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59917 
diff
changeset
 | 
679  | 
|
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
680  | 
fun make_context ths NONE cert =  | 
| 
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
681  | 
(Context.Theory (Context.certificate_theory cert)  | 
| 61051 | 682  | 
handle ERROR msg => raise CONTEXT (msg, [], [], ths, NONE))  | 
| 
61045
 
c7a7f063704a
clarified exceptions: avoid interference of formal context failure with regular rule application failure (which is routinely handled in user-space);
 
wenzelm 
parents: 
61044 
diff
changeset
 | 
683  | 
| make_context ths (SOME ctxt) cert =  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
684  | 
let  | 
| 
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
685  | 
val thy_id = Context.certificate_theory_id cert;  | 
| 
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
686  | 
val thy_id' = Context.theory_id (Proof_Context.theory_of ctxt);  | 
| 
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
687  | 
in  | 
| 
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
688  | 
if Context.subthy_id (thy_id, thy_id') then Context.Proof ctxt  | 
| 61051 | 689  | 
        else raise CONTEXT ("Bad context", [], [], ths, SOME (Context.Proof ctxt))
 | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
690  | 
end;  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
58946 
diff
changeset
 | 
691  | 
|
| 
64981
 
ea6199b23dfa
proper background certificate from make_context, which can be a super-theory of the direct join (amending d07464875dd4);
 
wenzelm 
parents: 
64574 
diff
changeset
 | 
692  | 
fun make_context_certificate ths opt_ctxt cert =  | 
| 
 
ea6199b23dfa
proper background certificate from make_context, which can be a super-theory of the direct join (amending d07464875dd4);
 
wenzelm 
parents: 
64574 
diff
changeset
 | 
693  | 
let  | 
| 
 
ea6199b23dfa
proper background certificate from make_context, which can be a super-theory of the direct join (amending d07464875dd4);
 
wenzelm 
parents: 
64574 
diff
changeset
 | 
694  | 
val context = make_context ths opt_ctxt cert;  | 
| 
 
ea6199b23dfa
proper background certificate from make_context, which can be a super-theory of the direct join (amending d07464875dd4);
 
wenzelm 
parents: 
64574 
diff
changeset
 | 
695  | 
val cert' = Context.Certificate (Context.theory_of context);  | 
| 
 
ea6199b23dfa
proper background certificate from make_context, which can be a super-theory of the direct join (amending d07464875dd4);
 
wenzelm 
parents: 
64574 
diff
changeset
 | 
696  | 
in (context, cert') end;  | 
| 
 
ea6199b23dfa
proper background certificate from make_context, which can be a super-theory of the direct join (amending d07464875dd4);
 
wenzelm 
parents: 
64574 
diff
changeset
 | 
697  | 
|
| 16945 | 698  | 
(*explicit weakening: maps |- B to A |- B*)  | 
699  | 
fun weaken raw_ct th =  | 
|
700  | 
let  | 
|
| 20261 | 701  | 
    val ct as Cterm {t = A, T, sorts, maxidx = maxidxA, ...} = adjust_maxidx_cterm ~1 raw_ct;
 | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
702  | 
    val Thm (der, {tags, maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = th;
 | 
| 16945 | 703  | 
in  | 
704  | 
if T <> propT then  | 
|
705  | 
      raise THM ("weaken: assumptions must have type prop", 0, [])
 | 
|
706  | 
else if maxidxA <> ~1 then  | 
|
707  | 
      raise THM ("weaken: assumptions may not contain schematic variables", maxidxA, [])
 | 
|
708  | 
else  | 
|
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
709  | 
Thm (der,  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
710  | 
       {cert = join_certificate1 (ct, th),
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
711  | 
tags = tags,  | 
| 16945 | 712  | 
maxidx = maxidx,  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
713  | 
constraints = constraints,  | 
| 16945 | 714  | 
shyps = Sorts.union sorts shyps,  | 
| 28354 | 715  | 
hyps = insert_hyps A hyps,  | 
| 16945 | 716  | 
tpairs = tpairs,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
717  | 
prop = prop})  | 
| 16945 | 718  | 
end;  | 
| 16656 | 719  | 
|
| 28624 | 720  | 
fun weaken_sorts raw_sorts ct =  | 
721  | 
let  | 
|
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
722  | 
    val Cterm {cert, t, T, maxidx, sorts} = ct;
 | 
| 61050 | 723  | 
val thy = theory_of_cterm ct;  | 
| 28624 | 724  | 
val more_sorts = Sorts.make (map (Sign.certify_sort thy) raw_sorts);  | 
725  | 
val sorts' = Sorts.union sorts more_sorts;  | 
|
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
726  | 
  in Cterm {cert = cert, t = t, T = T, maxidx = maxidx, sorts = sorts'} end;
 | 
| 28624 | 727  | 
|
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
728  | 
|
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
729  | 
|
| 32725 | 730  | 
(** derivations and promised proofs **)  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
731  | 
|
| 
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
 | 
732  | 
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
 | 
733  | 
  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
 | 
734  | 
|
| 70398 | 735  | 
val empty_deriv = make_deriv [] [] [] MinProof;  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
736  | 
|
| 28330 | 737  | 
|
| 28354 | 738  | 
(* inference rules *)  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
739  | 
|
| 70593 | 740  | 
val promise_ord: (serial * thm future) ord = fn ((i, _), (j, _)) => int_ord (j, i);  | 
| 28330 | 741  | 
|
| 
70835
 
2d991e01a671
proper treatment of axm_proof/oracle_proof like a closed proof constant, e.g. relevant for proof reconstruction of List.list.full_exhaustive_list.simps;
 
wenzelm 
parents: 
70834 
diff
changeset
 | 
742  | 
fun bad_proofs i =  | 
| 
 
2d991e01a671
proper treatment of axm_proof/oracle_proof like a closed proof constant, e.g. relevant for proof reconstruction of List.list.full_exhaustive_list.simps;
 
wenzelm 
parents: 
70834 
diff
changeset
 | 
743  | 
  error ("Illegal level of detail for proof objects: " ^ string_of_int i);
 | 
| 
 
2d991e01a671
proper treatment of axm_proof/oracle_proof like a closed proof constant, e.g. relevant for proof reconstruction of List.list.full_exhaustive_list.simps;
 
wenzelm 
parents: 
70834 
diff
changeset
 | 
744  | 
|
| 
52487
 
48bc24467008
backout dedd7952a62c: static "proofs" value within theory prevents later inferencing with different configuration;
 
wenzelm 
parents: 
52470 
diff
changeset
 | 
745  | 
fun deriv_rule2 f  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
746  | 
    (Deriv {promises = ps1, body = PBody {oracles = oracles1, thms = thms1, proof = prf1}})
 | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
747  | 
    (Deriv {promises = ps2, body = PBody {oracles = oracles2, thms = thms2, proof = prf2}}) =
 | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
748  | 
let  | 
| 39687 | 749  | 
val ps = Ord_List.union promise_ord ps1 ps2;  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
750  | 
val oracles = Proofterm.unions_oracles [oracles1, oracles2];  | 
| 44334 | 751  | 
val thms = Proofterm.unions_thms [thms1, thms2];  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
752  | 
val prf =  | 
| 
52487
 
48bc24467008
backout dedd7952a62c: static "proofs" value within theory prevents later inferencing with different configuration;
 
wenzelm 
parents: 
52470 
diff
changeset
 | 
753  | 
(case ! Proofterm.proofs of  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
754  | 
2 => f prf1 prf2  | 
| 
28804
 
5d3b1df16353
refined notion of derivation, consiting of promises and proof_body;
 
wenzelm 
parents: 
28675 
diff
changeset
 | 
755  | 
| 1 => MinProof  | 
| 
 
5d3b1df16353
refined notion of derivation, consiting of promises and proof_body;
 
wenzelm 
parents: 
28675 
diff
changeset
 | 
756  | 
| 0 => MinProof  | 
| 
70835
 
2d991e01a671
proper treatment of axm_proof/oracle_proof like a closed proof constant, e.g. relevant for proof reconstruction of List.list.full_exhaustive_list.simps;
 
wenzelm 
parents: 
70834 
diff
changeset
 | 
757  | 
| i => bad_proofs i);  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
758  | 
in make_deriv ps oracles thms prf end;  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
759  | 
|
| 
52487
 
48bc24467008
backout dedd7952a62c: static "proofs" value within theory prevents later inferencing with different configuration;
 
wenzelm 
parents: 
52470 
diff
changeset
 | 
760  | 
fun deriv_rule1 f = deriv_rule2 (K f) empty_deriv;  | 
| 
70812
 
3ae7949ef059
tuned -- allow slightly more expensive atomic proofs;
 
wenzelm 
parents: 
70603 
diff
changeset
 | 
761  | 
|
| 
 
3ae7949ef059
tuned -- allow slightly more expensive atomic proofs;
 
wenzelm 
parents: 
70603 
diff
changeset
 | 
762  | 
fun deriv_rule0 make_prf =  | 
| 
 
3ae7949ef059
tuned -- allow slightly more expensive atomic proofs;
 
wenzelm 
parents: 
70603 
diff
changeset
 | 
763  | 
if ! Proofterm.proofs <= 1 then empty_deriv  | 
| 
 
3ae7949ef059
tuned -- allow slightly more expensive atomic proofs;
 
wenzelm 
parents: 
70603 
diff
changeset
 | 
764  | 
else deriv_rule1 I (make_deriv [] [] [] (make_prf ()));  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
765  | 
|
| 
36621
 
2fd4e2c76636
proof terms for strip_shyps, based on the version by krauss/schropp with some notable differences:
 
wenzelm 
parents: 
36619 
diff
changeset
 | 
766  | 
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
 | 
767  | 
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
 | 
768  | 
|
| 1238 | 769  | 
|
| 32725 | 770  | 
(* fulfilled proofs *)  | 
771  | 
||
| 44331 | 772  | 
fun raw_promises_of (Thm (Deriv {promises, ...}, _)) = promises;
 | 
773  | 
||
774  | 
fun join_promises [] = ()  | 
|
775  | 
| join_promises promises = join_promises_of (Future.joins (map snd promises))  | 
|
| 
49008
 
a3cdb49c22cc
proper merge of promises to avoid exponential blow-up in pathologic situations (e.g. lack of PThm wrapping);
 
wenzelm 
parents: 
48263 
diff
changeset
 | 
776  | 
and join_promises_of thms = join_promises (Ord_List.make promise_ord (maps raw_promises_of thms));  | 
| 32725 | 777  | 
|
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
778  | 
fun fulfill_body (th as Thm (Deriv {promises, body}, _)) =
 | 
| 64571 | 779  | 
let val fulfilled_promises = map #1 promises ~~ map fulfill_body (Future.joins (map #2 promises))  | 
780  | 
in Proofterm.fulfill_norm_proof (theory_of_thm th) fulfilled_promises body end;  | 
|
| 32725 | 781  | 
|
| 
64574
 
1134e4d5e5b7
consolidate nested thms with persistent result, for improved performance;
 
wenzelm 
parents: 
64573 
diff
changeset
 | 
782  | 
fun proof_bodies_of thms = (join_promises_of thms; map fulfill_body thms);  | 
| 
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
 | 
783  | 
val proof_body_of = singleton proof_bodies_of;  | 
| 44331 | 784  | 
val proof_of = Proofterm.proof_of o proof_body_of;  | 
| 32725 | 785  | 
|
| 71088 | 786  | 
fun reconstruct_proof_of thm =  | 
787  | 
Proofterm.reconstruct_proof (theory_of_thm thm) (prop_of thm) (proof_of thm);  | 
|
788  | 
||
| 66168 | 789  | 
val consolidate = ignore o proof_bodies_of;  | 
| 32725 | 790  | 
|
| 71011 | 791  | 
fun expose_proofs thy thms =  | 
792  | 
if Proofterm.export_proof_boxes_required thy then  | 
|
| 
71018
 
d32ed8927a42
more robust expose_proofs corresponding to register_proofs/consolidate_theory;
 
wenzelm 
parents: 
71017 
diff
changeset
 | 
793  | 
Proofterm.export_proof_boxes (proof_bodies_of (map (transfer thy) thms))  | 
| 71011 | 794  | 
else ();  | 
795  | 
||
| 
71017
 
c85efa2be619
expose derivations more thoroughly, notably for locale/class reasoning;
 
wenzelm 
parents: 
71012 
diff
changeset
 | 
796  | 
fun expose_proof thy = expose_proofs thy o single;  | 
| 
 
c85efa2be619
expose derivations more thoroughly, notably for locale/class reasoning;
 
wenzelm 
parents: 
71012 
diff
changeset
 | 
797  | 
|
| 32725 | 798  | 
|
799  | 
(* future rule *)  | 
|
800  | 
||
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
801  | 
fun future_result i orig_cert orig_shyps orig_prop thm =  | 
| 32725 | 802  | 
let  | 
| 
36613
 
f3157c288aca
simplified primitive Thm.future: more direct theory check, do not hardwire strip_shyps here;
 
wenzelm 
parents: 
36330 
diff
changeset
 | 
803  | 
    fun err msg = raise THM ("future_result: " ^ msg, 0, [thm]);
 | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
804  | 
    val Thm (Deriv {promises, ...}, {cert, constraints, shyps, hyps, tpairs, prop, ...}) = thm;
 | 
| 
36613
 
f3157c288aca
simplified primitive Thm.future: more direct theory check, do not hardwire strip_shyps here;
 
wenzelm 
parents: 
36330 
diff
changeset
 | 
805  | 
|
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
806  | 
val _ = Context.eq_certificate (cert, orig_cert) orelse err "bad theory";  | 
| 32725 | 807  | 
val _ = prop aconv orig_prop orelse err "bad prop";  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
808  | 
val _ = null constraints orelse err "bad sort constraints";  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
809  | 
val _ = null tpairs orelse err "bad flex-flex constraints";  | 
| 32725 | 810  | 
val _ = null hyps orelse err "bad hyps";  | 
811  | 
val _ = Sorts.subset (shyps, orig_shyps) orelse err "bad shyps";  | 
|
812  | 
val _ = forall (fn (j, _) => i <> j) promises orelse err "bad dependencies";  | 
|
| 44331 | 813  | 
val _ = join_promises promises;  | 
| 32725 | 814  | 
in thm end;  | 
815  | 
||
816  | 
fun future future_thm ct =  | 
|
817  | 
let  | 
|
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
818  | 
    val Cterm {cert = cert, t = prop, T, maxidx, sorts} = ct;
 | 
| 32725 | 819  | 
    val _ = T <> propT andalso raise CTERM ("future: prop expected", [ct]);
 | 
| 70398 | 820  | 
val _ =  | 
821  | 
if Proofterm.proofs_enabled ()  | 
|
822  | 
      then raise CTERM ("future: proof terms enabled", [ct]) else ();
 | 
|
| 32725 | 823  | 
|
824  | 
val i = serial ();  | 
|
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
825  | 
val future = future_thm |> Future.map (future_result i cert sorts prop);  | 
| 32725 | 826  | 
in  | 
| 70398 | 827  | 
Thm (make_deriv [(i, future)] [] [] MinProof,  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
828  | 
     {cert = cert,
 | 
| 32725 | 829  | 
tags = [],  | 
830  | 
maxidx = maxidx,  | 
|
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
831  | 
constraints = [],  | 
| 32725 | 832  | 
shyps = sorts,  | 
833  | 
hyps = [],  | 
|
834  | 
tpairs = [],  | 
|
835  | 
prop = prop})  | 
|
836  | 
end;  | 
|
837  | 
||
838  | 
||
| 1238 | 839  | 
|
| 1529 | 840  | 
(** Axioms **)  | 
| 
387
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
841  | 
|
| 
70923
 
98d9b78b7f47
clarified axiom_table: uniform space (e.g. like consts), e.g. relevant for export of HOL-ex.Join_Theory;
 
wenzelm 
parents: 
70915 
diff
changeset
 | 
842  | 
fun axiom thy name =  | 
| 
 
98d9b78b7f47
clarified axiom_table: uniform space (e.g. like consts), e.g. relevant for export of HOL-ex.Join_Theory;
 
wenzelm 
parents: 
70915 
diff
changeset
 | 
843  | 
(case Name_Space.lookup (Theory.axiom_table thy) name of  | 
| 
 
98d9b78b7f47
clarified axiom_table: uniform space (e.g. like consts), e.g. relevant for export of HOL-ex.Join_Theory;
 
wenzelm 
parents: 
70915 
diff
changeset
 | 
844  | 
SOME prop =>  | 
| 
 
98d9b78b7f47
clarified axiom_table: uniform space (e.g. like consts), e.g. relevant for export of HOL-ex.Join_Theory;
 
wenzelm 
parents: 
70915 
diff
changeset
 | 
845  | 
let  | 
| 
 
98d9b78b7f47
clarified axiom_table: uniform space (e.g. like consts), e.g. relevant for export of HOL-ex.Join_Theory;
 
wenzelm 
parents: 
70915 
diff
changeset
 | 
846  | 
val der = deriv_rule0 (fn () => Proofterm.axm_proof name prop);  | 
| 
 
98d9b78b7f47
clarified axiom_table: uniform space (e.g. like consts), e.g. relevant for export of HOL-ex.Join_Theory;
 
wenzelm 
parents: 
70915 
diff
changeset
 | 
847  | 
val cert = Context.Certificate thy;  | 
| 
 
98d9b78b7f47
clarified axiom_table: uniform space (e.g. like consts), e.g. relevant for export of HOL-ex.Join_Theory;
 
wenzelm 
parents: 
70915 
diff
changeset
 | 
848  | 
val maxidx = maxidx_of_term prop;  | 
| 
 
98d9b78b7f47
clarified axiom_table: uniform space (e.g. like consts), e.g. relevant for export of HOL-ex.Join_Theory;
 
wenzelm 
parents: 
70915 
diff
changeset
 | 
849  | 
val shyps = Sorts.insert_term prop [];  | 
| 
 
98d9b78b7f47
clarified axiom_table: uniform space (e.g. like consts), e.g. relevant for export of HOL-ex.Join_Theory;
 
wenzelm 
parents: 
70915 
diff
changeset
 | 
850  | 
in  | 
| 
 
98d9b78b7f47
clarified axiom_table: uniform space (e.g. like consts), e.g. relevant for export of HOL-ex.Join_Theory;
 
wenzelm 
parents: 
70915 
diff
changeset
 | 
851  | 
Thm (der,  | 
| 
 
98d9b78b7f47
clarified axiom_table: uniform space (e.g. like consts), e.g. relevant for export of HOL-ex.Join_Theory;
 
wenzelm 
parents: 
70915 
diff
changeset
 | 
852  | 
          {cert = cert, tags = [], maxidx = maxidx,
 | 
| 
 
98d9b78b7f47
clarified axiom_table: uniform space (e.g. like consts), e.g. relevant for export of HOL-ex.Join_Theory;
 
wenzelm 
parents: 
70915 
diff
changeset
 | 
853  | 
constraints = [], shyps = shyps, hyps = [], tpairs = [], prop = prop})  | 
| 
 
98d9b78b7f47
clarified axiom_table: uniform space (e.g. like consts), e.g. relevant for export of HOL-ex.Join_Theory;
 
wenzelm 
parents: 
70915 
diff
changeset
 | 
854  | 
end  | 
| 
 
98d9b78b7f47
clarified axiom_table: uniform space (e.g. like consts), e.g. relevant for export of HOL-ex.Join_Theory;
 
wenzelm 
parents: 
70915 
diff
changeset
 | 
855  | 
  | NONE => raise THEORY ("No axiom " ^ quote name, [thy]));
 | 
| 
387
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
856  | 
|
| 
70923
 
98d9b78b7f47
clarified axiom_table: uniform space (e.g. like consts), e.g. relevant for export of HOL-ex.Join_Theory;
 
wenzelm 
parents: 
70915 
diff
changeset
 | 
857  | 
fun all_axioms_of thy =  | 
| 
 
98d9b78b7f47
clarified axiom_table: uniform space (e.g. like consts), e.g. relevant for export of HOL-ex.Join_Theory;
 
wenzelm 
parents: 
70915 
diff
changeset
 | 
858  | 
map (fn (name, _) => (name, axiom thy name)) (Theory.all_axioms_of thy);  | 
| 
776
 
df8f91c0e57c
improved axioms_of: returns thms as the manual says;
 
wenzelm 
parents: 
721 
diff
changeset
 | 
859  | 
|
| 6089 | 860  | 
|
| 
28804
 
5d3b1df16353
refined notion of derivation, consiting of promises and proof_body;
 
wenzelm 
parents: 
28675 
diff
changeset
 | 
861  | 
(* tags *)  | 
| 6089 | 862  | 
|
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
863  | 
val get_tags = #tags o rep_thm;  | 
| 6089 | 864  | 
|
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
865  | 
fun map_tags f (Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop})) =
 | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
866  | 
  Thm (der, {cert = cert, tags = f tags, maxidx = maxidx, constraints = constraints,
 | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
867  | 
shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop});  | 
| 0 | 868  | 
|
869  | 
||
| 
43795
 
ca5896a836ba
recovered some runtime sharing from d6b6c74a8bcf, without the global memory bottleneck;
 
wenzelm 
parents: 
43761 
diff
changeset
 | 
870  | 
(* technical adjustments *)  | 
| 
 
ca5896a836ba
recovered some runtime sharing from d6b6c74a8bcf, without the global memory bottleneck;
 
wenzelm 
parents: 
43761 
diff
changeset
 | 
871  | 
|
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
872  | 
fun norm_proof (th as Thm (der, args)) =  | 
| 
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
873  | 
Thm (deriv_rule1 (Proofterm.rew_proof (theory_of_thm th)) der, args);  | 
| 
23781
 
ab793a6ddf9f
Added function norm_proof for normalizing the proof term
 
berghofe 
parents: 
23657 
diff
changeset
 | 
874  | 
|
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
875  | 
fun adjust_maxidx_thm i  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
876  | 
    (th as Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop})) =
 | 
| 20261 | 877  | 
if maxidx = i then th  | 
878  | 
else if maxidx < i then  | 
|
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
879  | 
    Thm (der, {maxidx = i, cert = cert, tags = tags, constraints = constraints, shyps = shyps,
 | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
880  | 
hyps = hyps, tpairs = tpairs, prop = prop})  | 
| 20261 | 881  | 
else  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
882  | 
    Thm (der, {maxidx = Int.max (maxidx_tpairs tpairs (maxidx_of_term prop), i),
 | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
883  | 
cert = cert, tags = tags, constraints = constraints, shyps = shyps,  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
884  | 
hyps = hyps, tpairs = tpairs, prop = prop});  | 
| 564 | 885  | 
|
| 
387
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
886  | 
|
| 2509 | 887  | 
|
| 70454 | 888  | 
(*** Theory data ***)  | 
889  | 
||
| 70458 | 890  | 
(* type classes *)  | 
891  | 
||
| 70465 | 892  | 
structure Aritytab =  | 
893  | 
Table(  | 
|
894  | 
type key = string * sort list * class;  | 
|
895  | 
val ord =  | 
|
896  | 
fast_string_ord o apply2 #1  | 
|
| 73860 | 897  | 
||| fast_string_ord o apply2 #3  | 
898  | 
||| list_ord Term_Ord.sort_ord o apply2 #2;  | 
|
| 70465 | 899  | 
);  | 
900  | 
||
| 70458 | 901  | 
datatype classes = Classes of  | 
| 
70456
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
902  | 
 {classrels: thm Symreltab.table,
 | 
| 
70475
 
98b6da301e13
backed out changeset 1b8858f4c393: odd problems e.g. in CAVA_LTL_Modelchecker;
 
wenzelm 
parents: 
70472 
diff
changeset
 | 
903  | 
arities: (thm * string * serial) Aritytab.table};  | 
| 
70456
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
904  | 
|
| 70458 | 905  | 
fun make_classes (classrels, arities) = Classes {classrels = classrels, arities = arities};
 | 
| 
70456
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
906  | 
|
| 70465 | 907  | 
val empty_classes = make_classes (Symreltab.empty, Aritytab.empty);  | 
| 
70456
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
908  | 
|
| 70458 | 909  | 
(*see Theory.at_begin hook for transitive closure of classrels and arity completion*)  | 
910  | 
fun merge_classes  | 
|
911  | 
   (Classes {classrels = classrels1, arities = arities1},
 | 
|
912  | 
    Classes {classrels = classrels2, arities = arities2}) =
 | 
|
| 
70456
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
913  | 
let  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
914  | 
val classrels' = Symreltab.merge (K true) (classrels1, classrels2);  | 
| 70465 | 915  | 
val arities' = Aritytab.merge (K true) (arities1, arities2);  | 
| 70458 | 916  | 
in make_classes (classrels', arities') end;  | 
| 
70456
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
917  | 
|
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
918  | 
|
| 70458 | 919  | 
(* data *)  | 
920  | 
||
| 70454 | 921  | 
structure Data = Theory_Data  | 
922  | 
(  | 
|
923  | 
type T =  | 
|
| 
70456
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
924  | 
unit Name_Space.table * (*oracles: authentic derivation names*)  | 
| 70458 | 925  | 
classes; (*type classes within the logic*)  | 
| 
70456
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
926  | 
|
| 74112 | 927  | 
val empty : T = (Name_Space.empty_table Markup.oracleN, empty_classes);  | 
| 
70456
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
928  | 
fun merge ((oracles1, sorts1), (oracles2, sorts2)) : T =  | 
| 70458 | 929  | 
(Name_Space.merge_tables (oracles1, oracles2), merge_classes (sorts1, sorts2));  | 
| 70454 | 930  | 
);  | 
931  | 
||
| 
70456
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
932  | 
val get_oracles = #1 o Data.get;  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
933  | 
val map_oracles = Data.map o apfst;  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
934  | 
|
| 70458 | 935  | 
val get_classes = (fn (_, Classes args) => args) o Data.get;  | 
936  | 
val get_classrels = #classrels o get_classes;  | 
|
937  | 
val get_arities = #arities o get_classes;  | 
|
938  | 
||
939  | 
fun map_classes f =  | 
|
940  | 
  (Data.map o apsnd) (fn Classes {classrels, arities} => make_classes (f (classrels, arities)));
 | 
|
941  | 
fun map_classrels f = map_classes (fn (classrels, arities) => (f classrels, arities));  | 
|
942  | 
fun map_arities f = map_classes (fn (classrels, arities) => (classrels, f arities));  | 
|
943  | 
||
944  | 
||
945  | 
(* type classes *)  | 
|
946  | 
||
947  | 
fun the_classrel thy (c1, c2) =  | 
|
948  | 
(case Symreltab.lookup (get_classrels thy) (c1, c2) of  | 
|
949  | 
SOME thm => transfer thy thm  | 
|
950  | 
  | NONE => error ("Unproven class relation " ^
 | 
|
951  | 
Syntax.string_of_classrel (Proof_Context.init_global thy) [c1, c2]));  | 
|
952  | 
||
953  | 
fun the_arity thy (a, Ss, c) =  | 
|
| 70465 | 954  | 
(case Aritytab.lookup (get_arities thy) (a, Ss, c) of  | 
| 
70475
 
98b6da301e13
backed out changeset 1b8858f4c393: odd problems e.g. in CAVA_LTL_Modelchecker;
 
wenzelm 
parents: 
70472 
diff
changeset
 | 
955  | 
SOME (thm, _, _) => transfer thy thm  | 
| 70458 | 956  | 
  | NONE => error ("Unproven type arity " ^
 | 
957  | 
Syntax.string_of_arity (Proof_Context.init_global thy) (a, Ss, [c])));  | 
|
958  | 
||
959  | 
val classrel_proof = proof_of oo the_classrel;  | 
|
960  | 
val arity_proof = proof_of oo the_arity;  | 
|
| 70457 | 961  | 
|
| 70458 | 962  | 
|
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
963  | 
(* solve sort constraints by pro-forma proof *)  | 
| 70458 | 964  | 
|
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
965  | 
local  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
966  | 
|
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
967  | 
fun union_digest (oracles1, thms1) (oracles2, thms2) =  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
968  | 
(Proofterm.unions_oracles [oracles1, oracles2], Proofterm.unions_thms [thms1, thms2]);  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
969  | 
|
| 70587 | 970  | 
fun thm_digest (Thm (Deriv {body = PBody {oracles, thms, ...}, ...}, _)) =
 | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
971  | 
(oracles, thms);  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
972  | 
|
| 
70480
 
1a1b7d7f24bb
explicit check of left-over constraints from different theory, e.g. due to lack of Thm.trim_context;
 
wenzelm 
parents: 
70475 
diff
changeset
 | 
973  | 
fun constraint_digest ({theory = thy, typ, sort, ...}: constraint) =
 | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
974  | 
Sorts.of_sort_derivation (Sign.classes_of thy)  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
975  | 
   {class_relation = fn _ => fn _ => fn (digest, c1) => fn c2 =>
 | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
976  | 
if c1 = c2 then ([], []) else union_digest digest (thm_digest (the_classrel thy (c1, c2))),  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
977  | 
type_constructor = fn (a, _) => fn dom => fn c =>  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
978  | 
let val arity_digest = thm_digest (the_arity thy (a, (map o map) #2 dom, c))  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
979  | 
in (fold o fold) (union_digest o #1) dom arity_digest end,  | 
| 
70480
 
1a1b7d7f24bb
explicit check of left-over constraints from different theory, e.g. due to lack of Thm.trim_context;
 
wenzelm 
parents: 
70475 
diff
changeset
 | 
980  | 
type_variable = fn T => map (pair ([], [])) (Type.sort_of_atyp T)}  | 
| 
 
1a1b7d7f24bb
explicit check of left-over constraints from different theory, e.g. due to lack of Thm.trim_context;
 
wenzelm 
parents: 
70475 
diff
changeset
 | 
981  | 
(typ, sort);  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
982  | 
|
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
983  | 
in  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
984  | 
|
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
985  | 
fun solve_constraints (thm as Thm (_, {constraints = [], ...})) = thm
 | 
| 
70480
 
1a1b7d7f24bb
explicit check of left-over constraints from different theory, e.g. due to lack of Thm.trim_context;
 
wenzelm 
parents: 
70475 
diff
changeset
 | 
986  | 
| solve_constraints (thm as Thm (der, args)) =  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
987  | 
let  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
988  | 
        val {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop} = args;
 | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
989  | 
|
| 
70480
 
1a1b7d7f24bb
explicit check of left-over constraints from different theory, e.g. due to lack of Thm.trim_context;
 
wenzelm 
parents: 
70475 
diff
changeset
 | 
990  | 
val thy = Context.certificate_theory cert;  | 
| 
 
1a1b7d7f24bb
explicit check of left-over constraints from different theory, e.g. due to lack of Thm.trim_context;
 
wenzelm 
parents: 
70475 
diff
changeset
 | 
991  | 
val bad_thys =  | 
| 
 
1a1b7d7f24bb
explicit check of left-over constraints from different theory, e.g. due to lack of Thm.trim_context;
 
wenzelm 
parents: 
70475 
diff
changeset
 | 
992  | 
          constraints |> map_filter (fn {theory = thy', ...} =>
 | 
| 
 
1a1b7d7f24bb
explicit check of left-over constraints from different theory, e.g. due to lack of Thm.trim_context;
 
wenzelm 
parents: 
70475 
diff
changeset
 | 
993  | 
if Context.eq_thy (thy, thy') then NONE else SOME thy');  | 
| 
 
1a1b7d7f24bb
explicit check of left-over constraints from different theory, e.g. due to lack of Thm.trim_context;
 
wenzelm 
parents: 
70475 
diff
changeset
 | 
994  | 
val () =  | 
| 
 
1a1b7d7f24bb
explicit check of left-over constraints from different theory, e.g. due to lack of Thm.trim_context;
 
wenzelm 
parents: 
70475 
diff
changeset
 | 
995  | 
if null bad_thys then ()  | 
| 
 
1a1b7d7f24bb
explicit check of left-over constraints from different theory, e.g. due to lack of Thm.trim_context;
 
wenzelm 
parents: 
70475 
diff
changeset
 | 
996  | 
else  | 
| 
 
1a1b7d7f24bb
explicit check of left-over constraints from different theory, e.g. due to lack of Thm.trim_context;
 
wenzelm 
parents: 
70475 
diff
changeset
 | 
997  | 
            raise THEORY ("solve_constraints: bad theories for theorem\n" ^
 | 
| 
 
1a1b7d7f24bb
explicit check of left-over constraints from different theory, e.g. due to lack of Thm.trim_context;
 
wenzelm 
parents: 
70475 
diff
changeset
 | 
998  | 
Syntax.string_of_term_global thy (prop_of thm), thy :: bad_thys);  | 
| 
 
1a1b7d7f24bb
explicit check of left-over constraints from different theory, e.g. due to lack of Thm.trim_context;
 
wenzelm 
parents: 
70475 
diff
changeset
 | 
999  | 
|
| 70587 | 1000  | 
        val Deriv {promises, body = PBody {oracles, thms, proof}} = der;
 | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1001  | 
val (oracles', thms') = (oracles, thms)  | 
| 
70480
 
1a1b7d7f24bb
explicit check of left-over constraints from different theory, e.g. due to lack of Thm.trim_context;
 
wenzelm 
parents: 
70475 
diff
changeset
 | 
1002  | 
|> fold (fold union_digest o constraint_digest) constraints;  | 
| 70587 | 1003  | 
        val body' = PBody {oracles = oracles', thms = thms', proof = proof};
 | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1004  | 
in  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1005  | 
        Thm (Deriv {promises = promises, body = body'},
 | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1006  | 
          {constraints = [], cert = cert, tags = tags, maxidx = maxidx,
 | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1007  | 
shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop})  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1008  | 
end;  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1009  | 
|
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1010  | 
end;  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1011  | 
|
| 71527 | 1012  | 
(*Dangling sort constraints of a thm*)  | 
1013  | 
fun extra_shyps (th as Thm (_, {shyps, ...})) =
 | 
|
| 
74241
 
eb265f54e3ce
more efficient operations: traverse hyps only when required;
 
wenzelm 
parents: 
74235 
diff
changeset
 | 
1014  | 
  Sorts.subtract (fold_terms {hyps = true} Sorts.insert_term th []) shyps;
 | 
| 71527 | 1015  | 
|
1016  | 
(*Remove extra sorts that are witnessed by type signature information*)  | 
|
1017  | 
fun strip_shyps thm =  | 
|
1018  | 
(case thm of  | 
|
1019  | 
    Thm (_, {shyps = [], ...}) => thm
 | 
|
1020  | 
  | Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop}) =>
 | 
|
1021  | 
let  | 
|
1022  | 
val thy = theory_of_thm thm;  | 
|
1023  | 
||
1024  | 
val algebra = Sign.classes_of thy;  | 
|
1025  | 
val minimize = Sorts.minimize_sort algebra;  | 
|
1026  | 
val le = Sorts.sort_le algebra;  | 
|
1027  | 
fun lt (S1, S2) = le (S1, S2) andalso not (le (S2, S1));  | 
|
1028  | 
fun rel (S1, S2) = if S1 = S2 then [] else [(Term.aT S1, S2)];  | 
|
1029  | 
||
| 
74241
 
eb265f54e3ce
more efficient operations: traverse hyps only when required;
 
wenzelm 
parents: 
74235 
diff
changeset
 | 
1030  | 
val present =  | 
| 
 
eb265f54e3ce
more efficient operations: traverse hyps only when required;
 
wenzelm 
parents: 
74235 
diff
changeset
 | 
1031  | 
          (fold_terms {hyps = true} o fold_types o fold_atyps_sorts) (insert (eq_fst op =)) thm [];
 | 
| 71527 | 1032  | 
val extra = fold (Sorts.remove_sort o #2) present shyps;  | 
1033  | 
val witnessed = Sign.witness_sorts thy present extra;  | 
|
1034  | 
val non_witnessed = fold (Sorts.remove_sort o #2) witnessed extra |> map (`minimize);  | 
|
1035  | 
||
1036  | 
val extra' =  | 
|
1037  | 
non_witnessed |> map_filter (fn (S, _) =>  | 
|
1038  | 
if non_witnessed |> exists (fn (S', _) => lt (S', S)) then NONE else SOME S)  | 
|
1039  | 
|> Sorts.make;  | 
|
1040  | 
||
1041  | 
val constrs' =  | 
|
1042  | 
non_witnessed |> maps (fn (S1, S2) =>  | 
|
1043  | 
let val S0 = the (find_first (fn S => le (S, S1)) extra')  | 
|
1044  | 
in rel (S0, S1) @ rel (S1, S2) end);  | 
|
1045  | 
||
1046  | 
val constraints' = fold (insert_constraints thy) (witnessed @ constrs') constraints;  | 
|
1047  | 
val shyps' = fold (Sorts.insert_sort o #2) present extra';  | 
|
1048  | 
in  | 
|
1049  | 
Thm (deriv_rule_unconditional  | 
|
1050  | 
(Proofterm.strip_shyps_proof algebra present witnessed extra') der,  | 
|
1051  | 
         {cert = cert, tags = tags, maxidx = maxidx, constraints = constraints',
 | 
|
1052  | 
shyps = shyps', hyps = hyps, tpairs = tpairs, prop = prop})  | 
|
1053  | 
end)  | 
|
1054  | 
|> solve_constraints;  | 
|
1055  | 
||
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1056  | 
|
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1057  | 
|
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1058  | 
(*** Closed theorems with official name ***)  | 
| 70458 | 1059  | 
|
1060  | 
(*non-deterministic, depends on unknown promises*)  | 
|
1061  | 
fun derivation_closed (Thm (Deriv {body, ...}, _)) =
 | 
|
1062  | 
Proofterm.compact_proof (Proofterm.proof_of body);  | 
|
1063  | 
||
1064  | 
(*non-deterministic, depends on unknown promises*)  | 
|
| 
70543
 
33749040b6f8
clarified derivation_name vs. raw_derivation_name;
 
wenzelm 
parents: 
70517 
diff
changeset
 | 
1065  | 
fun raw_derivation_name (Thm (Deriv {body, ...}, {shyps, hyps, prop, ...})) =
 | 
| 70554 | 1066  | 
Proofterm.get_approximative_name shyps hyps prop (Proofterm.proof_of body);  | 
| 70458 | 1067  | 
|
| 
70915
 
bd4d37edfee4
clarified expand_proof/expand_name: allow more detailed control via thm_header;
 
wenzelm 
parents: 
70908 
diff
changeset
 | 
1068  | 
fun expand_name (Thm (Deriv {body, ...}, {shyps, hyps, prop, ...})) =
 | 
| 
 
bd4d37edfee4
clarified expand_proof/expand_name: allow more detailed control via thm_header;
 
wenzelm 
parents: 
70908 
diff
changeset
 | 
1069  | 
let  | 
| 
 
bd4d37edfee4
clarified expand_proof/expand_name: allow more detailed control via thm_header;
 
wenzelm 
parents: 
70908 
diff
changeset
 | 
1070  | 
val self_id =  | 
| 
 
bd4d37edfee4
clarified expand_proof/expand_name: allow more detailed control via thm_header;
 
wenzelm 
parents: 
70908 
diff
changeset
 | 
1071  | 
(case Proofterm.get_identity shyps hyps prop (Proofterm.proof_of body) of  | 
| 
 
bd4d37edfee4
clarified expand_proof/expand_name: allow more detailed control via thm_header;
 
wenzelm 
parents: 
70908 
diff
changeset
 | 
1072  | 
NONE => K false  | 
| 
 
bd4d37edfee4
clarified expand_proof/expand_name: allow more detailed control via thm_header;
 
wenzelm 
parents: 
70908 
diff
changeset
 | 
1073  | 
      | SOME {serial, ...} => fn (header: Proofterm.thm_header) => serial = #serial header);
 | 
| 
 
bd4d37edfee4
clarified expand_proof/expand_name: allow more detailed control via thm_header;
 
wenzelm 
parents: 
70908 
diff
changeset
 | 
1074  | 
fun expand header = if self_id header orelse #name header = "" then SOME "" else NONE;  | 
| 
 
bd4d37edfee4
clarified expand_proof/expand_name: allow more detailed control via thm_header;
 
wenzelm 
parents: 
70908 
diff
changeset
 | 
1075  | 
in expand end;  | 
| 
 
bd4d37edfee4
clarified expand_proof/expand_name: allow more detailed control via thm_header;
 
wenzelm 
parents: 
70908 
diff
changeset
 | 
1076  | 
|
| 
70543
 
33749040b6f8
clarified derivation_name vs. raw_derivation_name;
 
wenzelm 
parents: 
70517 
diff
changeset
 | 
1077  | 
(*deterministic name of finished proof*)  | 
| 
 
33749040b6f8
clarified derivation_name vs. raw_derivation_name;
 
wenzelm 
parents: 
70517 
diff
changeset
 | 
1078  | 
fun derivation_name (thm as Thm (_, {shyps, hyps, prop, ...})) =
 | 
| 70554 | 1079  | 
Proofterm.get_approximative_name shyps hyps prop (proof_of thm);  | 
| 
70543
 
33749040b6f8
clarified derivation_name vs. raw_derivation_name;
 
wenzelm 
parents: 
70517 
diff
changeset
 | 
1080  | 
|
| 
70595
 
2ae7e33c950f
clarified thm_id vs. thm_node/thm: retain theory_name;
 
wenzelm 
parents: 
70593 
diff
changeset
 | 
1081  | 
(*identified PThm node*)  | 
| 
70551
 
9e87e978be5e
clarified identity of PThm nodes: do not reuse old id after renaming -- enforce uniqueness of substructures;
 
wenzelm 
parents: 
70543 
diff
changeset
 | 
1082  | 
fun derivation_id (thm as Thm (_, {shyps, hyps, prop, ...})) =
 | 
| 70603 | 1083  | 
Proofterm.get_id shyps hyps prop (proof_of thm);  | 
| 
70551
 
9e87e978be5e
clarified identity of PThm nodes: do not reuse old id after renaming -- enforce uniqueness of substructures;
 
wenzelm 
parents: 
70543 
diff
changeset
 | 
1084  | 
|
| 
70595
 
2ae7e33c950f
clarified thm_id vs. thm_node/thm: retain theory_name;
 
wenzelm 
parents: 
70593 
diff
changeset
 | 
1085  | 
(*dependencies of PThm node*)  | 
| 
 
2ae7e33c950f
clarified thm_id vs. thm_node/thm: retain theory_name;
 
wenzelm 
parents: 
70593 
diff
changeset
 | 
1086  | 
fun thm_deps (thm as Thm (Deriv {promises = [], body = PBody {thms, ...}, ...}, _)) =
 | 
| 
 
2ae7e33c950f
clarified thm_id vs. thm_node/thm: retain theory_name;
 
wenzelm 
parents: 
70593 
diff
changeset
 | 
1087  | 
(case (derivation_id thm, thms) of  | 
| 
 
2ae7e33c950f
clarified thm_id vs. thm_node/thm: retain theory_name;
 
wenzelm 
parents: 
70593 
diff
changeset
 | 
1088  | 
        (SOME {serial = i, ...}, [(j, thm_node)]) =>
 | 
| 
 
2ae7e33c950f
clarified thm_id vs. thm_node/thm: retain theory_name;
 
wenzelm 
parents: 
70593 
diff
changeset
 | 
1089  | 
if i = j then Proofterm.thm_node_thms thm_node else thms  | 
| 
 
2ae7e33c950f
clarified thm_id vs. thm_node/thm: retain theory_name;
 
wenzelm 
parents: 
70593 
diff
changeset
 | 
1090  | 
| _ => thms)  | 
| 
 
2ae7e33c950f
clarified thm_id vs. thm_node/thm: retain theory_name;
 
wenzelm 
parents: 
70593 
diff
changeset
 | 
1091  | 
  | thm_deps thm = raise THM ("thm_deps: bad promises", 0, [thm]);
 | 
| 
 
2ae7e33c950f
clarified thm_id vs. thm_node/thm: retain theory_name;
 
wenzelm 
parents: 
70593 
diff
changeset
 | 
1092  | 
|
| 70494 | 1093  | 
fun name_derivation name_pos =  | 
| 
71530
 
046cf49162a3
more thorough strip_shyps for proof boxes (but types are usually stripped and reconstructed later);
 
wenzelm 
parents: 
71527 
diff
changeset
 | 
1094  | 
strip_shyps #> (fn thm as Thm (der, args) =>  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1095  | 
let  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1096  | 
val thy = theory_of_thm thm;  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1097  | 
|
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1098  | 
      val Deriv {promises, body} = der;
 | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1099  | 
      val {shyps, hyps, prop, tpairs, ...} = args;
 | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1100  | 
|
| 
70517
 
9a9003b5c0c1
more robust -- notably for metis, which tends to accumulate tpairs;
 
wenzelm 
parents: 
70503 
diff
changeset
 | 
1101  | 
      val _ = null tpairs orelse raise THM ("name_derivation: bad flex-flex constraints", 0, [thm]);
 | 
| 70458 | 1102  | 
|
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1103  | 
val ps = map (apsnd (Future.map fulfill_body)) promises;  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1104  | 
val (pthm, proof) =  | 
| 70494 | 1105  | 
Proofterm.thm_proof thy (classrel_proof thy) (arity_proof thy)  | 
1106  | 
name_pos shyps hyps prop ps body;  | 
|
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1107  | 
val der' = make_deriv [] [] [pthm] proof;  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1108  | 
in Thm (der', args) end);  | 
| 70458 | 1109  | 
|
| 70494 | 1110  | 
fun close_derivation pos =  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1111  | 
solve_constraints #> (fn thm =>  | 
| 
70517
 
9a9003b5c0c1
more robust -- notably for metis, which tends to accumulate tpairs;
 
wenzelm 
parents: 
70503 
diff
changeset
 | 
1112  | 
if not (null (tpairs_of thm)) orelse derivation_closed thm then thm  | 
| 
 
9a9003b5c0c1
more robust -- notably for metis, which tends to accumulate tpairs;
 
wenzelm 
parents: 
70503 
diff
changeset
 | 
1113  | 
    else name_derivation ("", pos) thm);
 | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1114  | 
|
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1115  | 
val trim_context = solve_constraints #> trim_context_thm;  | 
| 70454 | 1116  | 
|
1117  | 
||
1118  | 
||
1119  | 
(*** Oracles ***)  | 
|
1120  | 
||
1121  | 
fun add_oracle (b, oracle_fn) thy =  | 
|
1122  | 
let  | 
|
1123  | 
val (name, oracles') = Name_Space.define (Context.Theory thy) true (b, ()) (get_oracles thy);  | 
|
1124  | 
val thy' = map_oracles (K oracles') thy;  | 
|
1125  | 
fun invoke_oracle arg =  | 
|
1126  | 
      let val Cterm {cert = cert2, t = prop, T, maxidx, sorts} = oracle_fn arg in
 | 
|
1127  | 
if T <> propT then  | 
|
1128  | 
          raise THM ("Oracle's result must have type prop: " ^ name, 0, [])
 | 
|
1129  | 
else  | 
|
| 70834 | 1130  | 
let  | 
| 
70835
 
2d991e01a671
proper treatment of axm_proof/oracle_proof like a closed proof constant, e.g. relevant for proof reconstruction of List.list.full_exhaustive_list.simps;
 
wenzelm 
parents: 
70834 
diff
changeset
 | 
1131  | 
val (oracle, prf) =  | 
| 
 
2d991e01a671
proper treatment of axm_proof/oracle_proof like a closed proof constant, e.g. relevant for proof reconstruction of List.list.full_exhaustive_list.simps;
 
wenzelm 
parents: 
70834 
diff
changeset
 | 
1132  | 
(case ! Proofterm.proofs of  | 
| 
71465
 
910a081cca74
more position information for oracles (e.g. "skip_proof" for 'sorry'), requires Proofterm.proofs := 1;
 
wenzelm 
parents: 
71454 
diff
changeset
 | 
1133  | 
2 => (((name, Position.thread_data ()), SOME prop), Proofterm.oracle_proof name prop)  | 
| 
 
910a081cca74
more position information for oracles (e.g. "skip_proof" for 'sorry'), requires Proofterm.proofs := 1;
 
wenzelm 
parents: 
71454 
diff
changeset
 | 
1134  | 
| 1 => (((name, Position.thread_data ()), SOME prop), MinProof)  | 
| 
 
910a081cca74
more position information for oracles (e.g. "skip_proof" for 'sorry'), requires Proofterm.proofs := 1;
 
wenzelm 
parents: 
71454 
diff
changeset
 | 
1135  | 
| 0 => (((name, Position.none), NONE), MinProof)  | 
| 
70835
 
2d991e01a671
proper treatment of axm_proof/oracle_proof like a closed proof constant, e.g. relevant for proof reconstruction of List.list.full_exhaustive_list.simps;
 
wenzelm 
parents: 
70834 
diff
changeset
 | 
1136  | 
| i => bad_proofs i);  | 
| 70834 | 1137  | 
in  | 
| 70454 | 1138  | 
Thm (make_deriv [] [oracle] [] prf,  | 
1139  | 
             {cert = Context.join_certificate (Context.Certificate thy', cert2),
 | 
|
1140  | 
tags = [],  | 
|
1141  | 
maxidx = maxidx,  | 
|
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1142  | 
constraints = [],  | 
| 70454 | 1143  | 
shyps = sorts,  | 
1144  | 
hyps = [],  | 
|
1145  | 
tpairs = [],  | 
|
1146  | 
prop = prop})  | 
|
1147  | 
end  | 
|
1148  | 
end;  | 
|
1149  | 
in ((name, invoke_oracle), thy') end;  | 
|
1150  | 
||
| 70560 | 1151  | 
val oracle_space = Name_Space.space_of_table o get_oracles;  | 
1152  | 
||
1153  | 
fun pretty_oracle ctxt =  | 
|
1154  | 
Name_Space.pretty ctxt (oracle_space (Proof_Context.theory_of ctxt));  | 
|
1155  | 
||
| 70454 | 1156  | 
fun extern_oracles verbose ctxt =  | 
1157  | 
map #1 (Name_Space.markup_table verbose ctxt (get_oracles (Proof_Context.theory_of ctxt)));  | 
|
1158  | 
||
| 70565 | 1159  | 
fun check_oracle ctxt =  | 
1160  | 
Name_Space.check (Context.Proof ctxt) (get_oracles (Proof_Context.theory_of ctxt)) #> #1;  | 
|
1161  | 
||
| 70454 | 1162  | 
|
1163  | 
||
| 1529 | 1164  | 
(*** Meta rules ***)  | 
| 0 | 1165  | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1166  | 
(** primitive rules **)  | 
| 0 | 1167  | 
|
| 16656 | 1168  | 
(*The assumption rule A |- A*)  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1169  | 
fun assume raw_ct =  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
1170  | 
  let val Cterm {cert, 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
 | 
1171  | 
if T <> propT then  | 
| 19230 | 1172  | 
      raise THM ("assume: prop", 0, [])
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1173  | 
else if maxidx <> ~1 then  | 
| 19230 | 1174  | 
      raise THM ("assume: variables", maxidx, [])
 | 
| 
70812
 
3ae7949ef059
tuned -- allow slightly more expensive atomic proofs;
 
wenzelm 
parents: 
70603 
diff
changeset
 | 
1175  | 
else Thm (deriv_rule0 (fn () => Proofterm.Hyp prop),  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
1176  | 
     {cert = cert,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
1177  | 
tags = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1178  | 
maxidx = ~1,  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1179  | 
constraints = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1180  | 
shyps = sorts,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1181  | 
hyps = [prop],  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1182  | 
tpairs = [],  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1183  | 
prop = prop})  | 
| 0 | 1184  | 
end;  | 
1185  | 
||
| 1220 | 1186  | 
(*Implication introduction  | 
| 3529 | 1187  | 
[A]  | 
1188  | 
:  | 
|
1189  | 
B  | 
|
| 1220 | 1190  | 
-------  | 
| 67721 | 1191  | 
A \<Longrightarrow> B  | 
| 1220 | 1192  | 
*)  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1193  | 
fun implies_intr  | 
| 68691 | 1194  | 
    (ct as Cterm {t = A, T, maxidx = maxidx1, sorts, ...})
 | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1195  | 
    (th as Thm (der, {maxidx = maxidx2, hyps, constraints, shyps, tpairs, prop, ...})) =
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1196  | 
if T <> propT then  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1197  | 
    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
 | 
1198  | 
else  | 
| 
52487
 
48bc24467008
backout dedd7952a62c: static "proofs" value within theory prevents later inferencing with different configuration;
 
wenzelm 
parents: 
52470 
diff
changeset
 | 
1199  | 
Thm (deriv_rule1 (Proofterm.implies_intr_proof A) der,  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
1200  | 
     {cert = join_certificate1 (ct, th),
 | 
| 
52487
 
48bc24467008
backout dedd7952a62c: static "proofs" value within theory prevents later inferencing with different configuration;
 
wenzelm 
parents: 
52470 
diff
changeset
 | 
1201  | 
tags = [],  | 
| 68691 | 1202  | 
maxidx = Int.max (maxidx1, maxidx2),  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1203  | 
constraints = constraints,  | 
| 
52487
 
48bc24467008
backout dedd7952a62c: static "proofs" value within theory prevents later inferencing with different configuration;
 
wenzelm 
parents: 
52470 
diff
changeset
 | 
1204  | 
shyps = Sorts.union sorts shyps,  | 
| 
 
48bc24467008
backout dedd7952a62c: static "proofs" value within theory prevents later inferencing with different configuration;
 
wenzelm 
parents: 
52470 
diff
changeset
 | 
1205  | 
hyps = remove_hyps A hyps,  | 
| 
 
48bc24467008
backout dedd7952a62c: static "proofs" value within theory prevents later inferencing with different configuration;
 
wenzelm 
parents: 
52470 
diff
changeset
 | 
1206  | 
tpairs = tpairs,  | 
| 
 
48bc24467008
backout dedd7952a62c: static "proofs" value within theory prevents later inferencing with different configuration;
 
wenzelm 
parents: 
52470 
diff
changeset
 | 
1207  | 
prop = Logic.mk_implies (A, prop)});  | 
| 0 | 1208  | 
|
| 1529 | 1209  | 
|
| 1220 | 1210  | 
(*Implication elimination  | 
| 67721 | 1211  | 
A \<Longrightarrow> B A  | 
| 1220 | 1212  | 
------------  | 
1213  | 
B  | 
|
1214  | 
*)  | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1215  | 
fun implies_elim thAB thA =  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1216  | 
let  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1217  | 
val Thm (derA,  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1218  | 
      {maxidx = maxidx1, hyps = hypsA, constraints = constraintsA, shyps = shypsA,
 | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1219  | 
tpairs = tpairsA, prop = propA, ...}) = thA  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1220  | 
    and Thm (der, {maxidx = maxidx2, hyps, constraints, shyps, tpairs, prop, ...}) = thAB;
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1221  | 
    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
 | 
1222  | 
in  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1223  | 
(case prop of  | 
| 56245 | 1224  | 
      Const ("Pure.imp", _) $ A $ B =>
 | 
| 20512 | 1225  | 
if A aconv propA then  | 
| 
52487
 
48bc24467008
backout dedd7952a62c: static "proofs" value within theory prevents later inferencing with different configuration;
 
wenzelm 
parents: 
52470 
diff
changeset
 | 
1226  | 
Thm (deriv_rule2 (curry Proofterm.%%) der derA,  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
1227  | 
           {cert = join_certificate2 (thAB, thA),
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
1228  | 
tags = [],  | 
| 68691 | 1229  | 
maxidx = Int.max (maxidx1, maxidx2),  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1230  | 
constraints = union_constraints constraintsA constraints,  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1231  | 
shyps = Sorts.union shypsA shyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1232  | 
hyps = union_hyps hypsA hyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1233  | 
tpairs = union_tpairs tpairsA tpairs,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1234  | 
prop = B})  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1235  | 
else err ()  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1236  | 
| _ => err ())  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1237  | 
end;  | 
| 250 | 1238  | 
|
| 1220 | 1239  | 
(*Forall introduction. The Free or Var x must not be free in the hypotheses.  | 
| 16656 | 1240  | 
[x]  | 
1241  | 
:  | 
|
1242  | 
A  | 
|
1243  | 
------  | 
|
| 67721 | 1244  | 
\<And>x. A  | 
| 1220 | 1245  | 
*)  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1246  | 
fun forall_intr  | 
| 
68692
 
0c568ec56f37
proper maxidx: if x does not occur in A, its maxidx could get lost;
 
wenzelm 
parents: 
68691 
diff
changeset
 | 
1247  | 
    (ct as Cterm {maxidx = maxidx1, t = x, T, sorts, ...})
 | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1248  | 
    (th as Thm (der, {maxidx = maxidx2, constraints, shyps, hyps, tpairs, prop, ...})) =
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1249  | 
let  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1250  | 
fun result a =  | 
| 70814 | 1251  | 
Thm (deriv_rule1 (Proofterm.forall_intr_proof (a, x) NONE) der,  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
1252  | 
       {cert = join_certificate1 (ct, th),
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
1253  | 
tags = [],  | 
| 
68692
 
0c568ec56f37
proper maxidx: if x does not occur in A, its maxidx could get lost;
 
wenzelm 
parents: 
68691 
diff
changeset
 | 
1254  | 
maxidx = Int.max (maxidx1, maxidx2),  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1255  | 
constraints = constraints,  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1256  | 
shyps = Sorts.union sorts shyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1257  | 
hyps = hyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1258  | 
tpairs = tpairs,  | 
| 
46217
 
7b19666f0e3d
renamed Term.all to Logic.all_const, in accordance to HOLogic.all_const;
 
wenzelm 
parents: 
45443 
diff
changeset
 | 
1259  | 
prop = Logic.all_const T $ Abs (a, T, abstract_over (x, prop))});  | 
| 21798 | 1260  | 
fun check_occs a x ts =  | 
| 
16847
 
8fc160b12e73
invoke_oracle: do not keep theory value, but theory_ref;
 
wenzelm 
parents: 
16725 
diff
changeset
 | 
1261  | 
if exists (fn t => Logic.occs (x, t)) ts then  | 
| 21798 | 1262  | 
        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
 | 
1263  | 
else ();  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1264  | 
in  | 
| 52788 | 1265  | 
(case x of  | 
| 21798 | 1266  | 
Free (a, _) => (check_occs a x hyps; check_occs a x (terms_of_tpairs tpairs); result a)  | 
1267  | 
| Var ((a, _), _) => (check_occs a x (terms_of_tpairs tpairs); result a)  | 
|
| 52788 | 1268  | 
    | _ => raise THM ("forall_intr: not a variable", 0, [th]))
 | 
| 0 | 1269  | 
end;  | 
1270  | 
||
| 1220 | 1271  | 
(*Forall elimination  | 
| 67721 | 1272  | 
\<And>x. A  | 
| 1220 | 1273  | 
------  | 
1274  | 
A[t/x]  | 
|
1275  | 
*)  | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1276  | 
fun forall_elim  | 
| 68691 | 1277  | 
    (ct as Cterm {t, T, maxidx = maxidx1, sorts, ...})
 | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1278  | 
    (th as Thm (der, {maxidx = maxidx2, constraints, shyps, hyps, tpairs, prop, ...})) =
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1279  | 
(case prop of  | 
| 56245 | 1280  | 
    Const ("Pure.all", Type ("fun", [Type ("fun", [qary, _]), _])) $ A =>
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1281  | 
if T <> qary then  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1282  | 
        raise THM ("forall_elim: type mismatch", 0, [th])
 | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1283  | 
else  | 
| 
52487
 
48bc24467008
backout dedd7952a62c: static "proofs" value within theory prevents later inferencing with different configuration;
 
wenzelm 
parents: 
52470 
diff
changeset
 | 
1284  | 
Thm (deriv_rule1 (Proofterm.% o rpair (SOME t)) der,  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
1285  | 
         {cert = join_certificate1 (ct, th),
 | 
| 
52487
 
48bc24467008
backout dedd7952a62c: static "proofs" value within theory prevents later inferencing with different configuration;
 
wenzelm 
parents: 
52470 
diff
changeset
 | 
1286  | 
tags = [],  | 
| 68691 | 1287  | 
maxidx = Int.max (maxidx1, maxidx2),  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1288  | 
constraints = constraints,  | 
| 
52487
 
48bc24467008
backout dedd7952a62c: static "proofs" value within theory prevents later inferencing with different configuration;
 
wenzelm 
parents: 
52470 
diff
changeset
 | 
1289  | 
shyps = Sorts.union sorts shyps,  | 
| 
 
48bc24467008
backout dedd7952a62c: static "proofs" value within theory prevents later inferencing with different configuration;
 
wenzelm 
parents: 
52470 
diff
changeset
 | 
1290  | 
hyps = hyps,  | 
| 
 
48bc24467008
backout dedd7952a62c: static "proofs" value within theory prevents later inferencing with different configuration;
 
wenzelm 
parents: 
52470 
diff
changeset
 | 
1291  | 
tpairs = tpairs,  | 
| 
 
48bc24467008
backout dedd7952a62c: static "proofs" value within theory prevents later inferencing with different configuration;
 
wenzelm 
parents: 
52470 
diff
changeset
 | 
1292  | 
prop = Term.betapply (A, t)})  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1293  | 
  | _ => raise THM ("forall_elim: not quantified", 0, [th]));
 | 
| 0 | 1294  | 
|
1295  | 
||
| 1220 | 1296  | 
(* Equality *)  | 
| 0 | 1297  | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1298  | 
(*Reflexivity  | 
| 67721 | 1299  | 
t \<equiv> t  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1300  | 
*)  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
1301  | 
fun reflexive (Cterm {cert, t, T = _, maxidx, sorts}) =
 | 
| 70814 | 1302  | 
Thm (deriv_rule0 (fn () => Proofterm.reflexive_proof),  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
1303  | 
   {cert = cert,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
1304  | 
tags = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1305  | 
maxidx = maxidx,  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1306  | 
constraints = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1307  | 
shyps = sorts,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1308  | 
hyps = [],  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1309  | 
tpairs = [],  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1310  | 
prop = Logic.mk_equals (t, t)});  | 
| 0 | 1311  | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1312  | 
(*Symmetry  | 
| 67721 | 1313  | 
t \<equiv> u  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1314  | 
------  | 
| 67721 | 1315  | 
u \<equiv> t  | 
| 1220 | 1316  | 
*)  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1317  | 
fun symmetric (th as Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...})) =
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1318  | 
(case prop of  | 
| 56245 | 1319  | 
    (eq as Const ("Pure.eq", _)) $ t $ u =>
 | 
| 70814 | 1320  | 
Thm (deriv_rule1 Proofterm.symmetric_proof der,  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
1321  | 
       {cert = cert,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
1322  | 
tags = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1323  | 
maxidx = maxidx,  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1324  | 
constraints = constraints,  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1325  | 
shyps = shyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1326  | 
hyps = hyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1327  | 
tpairs = tpairs,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1328  | 
prop = eq $ u $ t})  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1329  | 
    | _ => raise THM ("symmetric", 0, [th]));
 | 
| 0 | 1330  | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1331  | 
(*Transitivity  | 
| 67721 | 1332  | 
t1 \<equiv> u u \<equiv> t2  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1333  | 
------------------  | 
| 67721 | 1334  | 
t1 \<equiv> t2  | 
| 1220 | 1335  | 
*)  | 
| 0 | 1336  | 
fun transitive th1 th2 =  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1337  | 
let  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1338  | 
    val Thm (der1, {maxidx = maxidx1, hyps = hyps1, constraints = constraints1, shyps = shyps1,
 | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1339  | 
tpairs = tpairs1, prop = prop1, ...}) = th1  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1340  | 
    and Thm (der2, {maxidx = maxidx2, hyps = hyps2, constraints = constraints2, shyps = shyps2,
 | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1341  | 
tpairs = tpairs2, prop = prop2, ...}) = th2;  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1342  | 
    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
 | 
1343  | 
in  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1344  | 
case (prop1, prop2) of  | 
| 70814 | 1345  | 
      ((eq as Const ("Pure.eq", Type (_, [U, _]))) $ t1 $ u, Const ("Pure.eq", _) $ u' $ t2) =>
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1346  | 
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
 | 
1347  | 
else  | 
| 70814 | 1348  | 
Thm (deriv_rule2 (Proofterm.transitive_proof U u) der1 der2,  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
1349  | 
           {cert = join_certificate2 (th1, th2),
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
1350  | 
tags = [],  | 
| 68691 | 1351  | 
maxidx = Int.max (maxidx1, maxidx2),  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1352  | 
constraints = union_constraints constraints1 constraints2,  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1353  | 
shyps = Sorts.union shyps1 shyps2,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1354  | 
hyps = union_hyps hyps1 hyps2,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1355  | 
tpairs = union_tpairs tpairs1 tpairs2,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1356  | 
prop = eq $ t1 $ t2})  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1357  | 
| _ => err "premises"  | 
| 0 | 1358  | 
end;  | 
1359  | 
||
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1360  | 
(*Beta-conversion  | 
| 67721 | 1361  | 
(\<lambda>x. t) u \<equiv> t[u/x]  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1362  | 
fully beta-reduces the term if full = true  | 
| 
10416
 
5b33e732e459
- Moved rewriting functions to meta_simplifier.ML
 
berghofe 
parents: 
10348 
diff
changeset
 | 
1363  | 
*)  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
1364  | 
fun beta_conversion full (Cterm {cert, t, T = _, maxidx, sorts}) =
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1365  | 
let val t' =  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1366  | 
if full then Envir.beta_norm t  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1367  | 
else  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1368  | 
(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
 | 
1369  | 
      | _ => raise THM ("beta_conversion: not a redex", 0, []));
 | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1370  | 
in  | 
| 70814 | 1371  | 
Thm (deriv_rule0 (fn () => Proofterm.reflexive_proof),  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
1372  | 
     {cert = cert,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
1373  | 
tags = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1374  | 
maxidx = maxidx,  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1375  | 
constraints = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1376  | 
shyps = sorts,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1377  | 
hyps = [],  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1378  | 
tpairs = [],  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1379  | 
prop = Logic.mk_equals (t, t')})  | 
| 
10416
 
5b33e732e459
- Moved rewriting functions to meta_simplifier.ML
 
berghofe 
parents: 
10348 
diff
changeset
 | 
1380  | 
end;  | 
| 
 
5b33e732e459
- Moved rewriting functions to meta_simplifier.ML
 
berghofe 
parents: 
10348 
diff
changeset
 | 
1381  | 
|
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
1382  | 
fun eta_conversion (Cterm {cert, t, T = _, maxidx, sorts}) =
 | 
| 70814 | 1383  | 
Thm (deriv_rule0 (fn () => Proofterm.reflexive_proof),  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
1384  | 
   {cert = cert,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
1385  | 
tags = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1386  | 
maxidx = maxidx,  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1387  | 
constraints = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1388  | 
shyps = sorts,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1389  | 
hyps = [],  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1390  | 
tpairs = [],  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1391  | 
prop = Logic.mk_equals (t, Envir.eta_contract t)});  | 
| 0 | 1392  | 
|
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
1393  | 
fun eta_long_conversion (Cterm {cert, t, T = _, maxidx, sorts}) =
 | 
| 70814 | 1394  | 
Thm (deriv_rule0 (fn () => Proofterm.reflexive_proof),  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
1395  | 
   {cert = cert,
 | 
| 23493 | 1396  | 
tags = [],  | 
1397  | 
maxidx = maxidx,  | 
|
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1398  | 
constraints = [],  | 
| 23493 | 1399  | 
shyps = sorts,  | 
1400  | 
hyps = [],  | 
|
1401  | 
tpairs = [],  | 
|
| 52131 | 1402  | 
prop = Logic.mk_equals (t, Envir.eta_long [] t)});  | 
| 23493 | 1403  | 
|
| 0 | 1404  | 
(*The abstraction rule. The Free or Var x must not be free in the hypotheses.  | 
1405  | 
The bound variable will be named "a" (since x will be something like x320)  | 
|
| 67721 | 1406  | 
t \<equiv> u  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1407  | 
--------------  | 
| 67721 | 1408  | 
\<lambda>x. t \<equiv> \<lambda>x. u  | 
| 1220 | 1409  | 
*)  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1410  | 
fun abstract_rule a  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1411  | 
    (Cterm {t = x, T, sorts, ...})
 | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1412  | 
    (th as Thm (der, {cert, maxidx, hyps, constraints, shyps, tpairs, prop, ...})) =
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1413  | 
let  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1414  | 
val (t, u) = Logic.dest_equals prop  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1415  | 
      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
 | 
1416  | 
val result =  | 
| 70814 | 1417  | 
Thm (deriv_rule1 (Proofterm.abstract_rule_proof (a, x)) der,  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
1418  | 
       {cert = cert,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
1419  | 
tags = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1420  | 
maxidx = maxidx,  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1421  | 
constraints = constraints,  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1422  | 
shyps = Sorts.union sorts shyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1423  | 
hyps = hyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1424  | 
tpairs = tpairs,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1425  | 
prop = Logic.mk_equals  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1426  | 
(Abs (a, T, abstract_over (x, t)), Abs (a, T, abstract_over (x, u)))});  | 
| 21798 | 1427  | 
fun check_occs a x ts =  | 
| 
16847
 
8fc160b12e73
invoke_oracle: do not keep theory value, but theory_ref;
 
wenzelm 
parents: 
16725 
diff
changeset
 | 
1428  | 
if exists (fn t => Logic.occs (x, t)) ts then  | 
| 21798 | 1429  | 
        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
 | 
1430  | 
else ();  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1431  | 
in  | 
| 52788 | 1432  | 
(case x of  | 
| 21798 | 1433  | 
Free (a, _) => (check_occs a x hyps; check_occs a x (terms_of_tpairs tpairs); result)  | 
1434  | 
| Var ((a, _), _) => (check_occs a x (terms_of_tpairs tpairs); result)  | 
|
| 52788 | 1435  | 
    | _ => raise THM ("abstract_rule: not a variable", 0, [th]))
 | 
| 0 | 1436  | 
end;  | 
1437  | 
||
1438  | 
(*The combination rule  | 
|
| 67721 | 1439  | 
f \<equiv> g t \<equiv> u  | 
1440  | 
-------------  | 
|
1441  | 
f t \<equiv> g u  | 
|
| 1220 | 1442  | 
*)  | 
| 0 | 1443  | 
fun combination th1 th2 =  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1444  | 
let  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1445  | 
    val Thm (der1, {maxidx = maxidx1, constraints = constraints1, shyps = shyps1,
 | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1446  | 
hyps = hyps1, tpairs = tpairs1, prop = prop1, ...}) = th1  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1447  | 
    and Thm (der2, {maxidx = maxidx2, constraints = constraints2, shyps = shyps2,
 | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1448  | 
hyps = hyps2, tpairs = tpairs2, prop = prop2, ...}) = th2;  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1449  | 
fun chktypes fT tT =  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1450  | 
(case fT of  | 
| 32784 | 1451  | 
        Type ("fun", [T1, _]) =>
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1452  | 
if T1 <> tT then  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1453  | 
            raise THM ("combination: types", 0, [th1, th2])
 | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1454  | 
else ()  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1455  | 
      | _ => 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
 | 
1456  | 
in  | 
| 52788 | 1457  | 
(case (prop1, prop2) of  | 
| 56245 | 1458  | 
      (Const ("Pure.eq", Type ("fun", [fT, _])) $ f $ g,
 | 
1459  | 
       Const ("Pure.eq", Type ("fun", [tT, _])) $ t $ u) =>
 | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1460  | 
(chktypes fT tT;  | 
| 70908 | 1461  | 
Thm (deriv_rule2 (Proofterm.combination_proof f g t u) der1 der2,  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
1462  | 
           {cert = join_certificate2 (th1, th2),
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
1463  | 
tags = [],  | 
| 68691 | 1464  | 
maxidx = Int.max (maxidx1, maxidx2),  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1465  | 
constraints = union_constraints constraints1 constraints2,  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1466  | 
shyps = Sorts.union shyps1 shyps2,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1467  | 
hyps = union_hyps hyps1 hyps2,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1468  | 
tpairs = union_tpairs tpairs1 tpairs2,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1469  | 
prop = Logic.mk_equals (f $ t, g $ u)}))  | 
| 52788 | 1470  | 
     | _ => raise THM ("combination: premises", 0, [th1, th2]))
 | 
| 0 | 1471  | 
end;  | 
1472  | 
||
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1473  | 
(*Equality introduction  | 
| 67721 | 1474  | 
A \<Longrightarrow> B B \<Longrightarrow> A  | 
| 3529 | 1475  | 
----------------  | 
| 67721 | 1476  | 
A \<equiv> B  | 
| 1220 | 1477  | 
*)  | 
| 0 | 1478  | 
fun equal_intr th1 th2 =  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1479  | 
let  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1480  | 
    val Thm (der1, {maxidx = maxidx1, constraints = constraints1, shyps = shyps1,
 | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1481  | 
hyps = hyps1, tpairs = tpairs1, prop = prop1, ...}) = th1  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1482  | 
    and Thm (der2, {maxidx = maxidx2, constraints = constraints2, shyps = shyps2,
 | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1483  | 
hyps = hyps2, tpairs = tpairs2, prop = prop2, ...}) = th2;  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1484  | 
    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
 | 
1485  | 
in  | 
| 52788 | 1486  | 
(case (prop1, prop2) of  | 
| 56245 | 1487  | 
      (Const("Pure.imp", _) $ A $ B, Const("Pure.imp", _) $ B' $ A') =>
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1488  | 
if A aconv A' andalso B aconv B' then  | 
| 70814 | 1489  | 
Thm (deriv_rule2 (Proofterm.equal_intr_proof A B) der1 der2,  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
1490  | 
           {cert = join_certificate2 (th1, th2),
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
1491  | 
tags = [],  | 
| 68691 | 1492  | 
maxidx = Int.max (maxidx1, maxidx2),  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1493  | 
constraints = union_constraints constraints1 constraints2,  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1494  | 
shyps = Sorts.union shyps1 shyps2,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1495  | 
hyps = union_hyps hyps1 hyps2,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1496  | 
tpairs = union_tpairs tpairs1 tpairs2,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1497  | 
prop = Logic.mk_equals (A, B)})  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1498  | 
else err "not equal"  | 
| 52788 | 1499  | 
| _ => err "premises")  | 
| 1529 | 1500  | 
end;  | 
1501  | 
||
1502  | 
(*The equal propositions rule  | 
|
| 67721 | 1503  | 
A \<equiv> B A  | 
| 1529 | 1504  | 
---------  | 
1505  | 
B  | 
|
1506  | 
*)  | 
|
1507  | 
fun equal_elim th1 th2 =  | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1508  | 
let  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1509  | 
    val Thm (der1, {maxidx = maxidx1, constraints = constraints1, shyps = shyps1,
 | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1510  | 
hyps = hyps1, tpairs = tpairs1, prop = prop1, ...}) = th1  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1511  | 
    and Thm (der2, {maxidx = maxidx2, constraints = constraints2, shyps = shyps2,
 | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1512  | 
hyps = hyps2, tpairs = tpairs2, prop = prop2, ...}) = th2;  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1513  | 
    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
 | 
1514  | 
in  | 
| 52788 | 1515  | 
(case prop1 of  | 
| 56245 | 1516  | 
      Const ("Pure.eq", _) $ A $ B =>
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1517  | 
if prop2 aconv A then  | 
| 70814 | 1518  | 
Thm (deriv_rule2 (Proofterm.equal_elim_proof A B) der1 der2,  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
1519  | 
           {cert = join_certificate2 (th1, th2),
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
1520  | 
tags = [],  | 
| 68691 | 1521  | 
maxidx = Int.max (maxidx1, maxidx2),  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1522  | 
constraints = union_constraints constraints1 constraints2,  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1523  | 
shyps = Sorts.union shyps1 shyps2,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1524  | 
hyps = union_hyps hyps1 hyps2,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1525  | 
tpairs = union_tpairs tpairs1 tpairs2,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1526  | 
prop = B})  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1527  | 
else err "not equal"  | 
| 52789 | 1528  | 
| _ => err "major premise")  | 
| 1529 | 1529  | 
end;  | 
| 0 | 1530  | 
|
| 1220 | 1531  | 
|
1532  | 
||
| 0 | 1533  | 
(**** Derived rules ****)  | 
1534  | 
||
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1535  | 
(*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
 | 
1536  | 
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
 | 
1537  | 
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
 | 
1538  | 
flex-flex.*)  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1539  | 
fun flexflex_rule opt_ctxt =  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1540  | 
solve_constraints #> (fn th =>  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1541  | 
let  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1542  | 
      val Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = th;
 | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1543  | 
val (context, cert') = make_context_certificate [th] opt_ctxt cert;  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1544  | 
in  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1545  | 
Unify.smash_unifiers context tpairs (Envir.empty maxidx)  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1546  | 
|> Seq.map (fn env =>  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1547  | 
if Envir.is_empty env then th  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1548  | 
else  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1549  | 
let  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1550  | 
val tpairs' = tpairs |> map (apply2 (Envir.norm_term env))  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1551  | 
(*remove trivial tpairs, of the form t \<equiv> t*)  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1552  | 
|> filter_out (op aconv);  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1553  | 
val der' = deriv_rule1 (Proofterm.norm_proof' env) der;  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1554  | 
val constraints' =  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1555  | 
insert_constraints_env (Context.certificate_theory cert') env constraints;  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1556  | 
val prop' = Envir.norm_term env prop;  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1557  | 
val maxidx = maxidx_tpairs tpairs' (maxidx_of_term prop');  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1558  | 
val shyps = Envir.insert_sorts env shyps;  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1559  | 
in  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1560  | 
              Thm (der', {cert = cert', tags = [], maxidx = maxidx, constraints = constraints',
 | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1561  | 
shyps = shyps, hyps = hyps, tpairs = tpairs', prop = prop'})  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1562  | 
end)  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1563  | 
end);  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1564  | 
|
| 0 | 1565  | 
|
| 19910 | 1566  | 
(*Generalization of fixed variables  | 
1567  | 
A  | 
|
1568  | 
--------------------  | 
|
1569  | 
A[?'a/'a, ?x/x, ...]  | 
|
1570  | 
*)  | 
|
1571  | 
||
| 
74200
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1572  | 
fun generalize (tfrees, frees) idx th =  | 
| 74266 | 1573  | 
if Names.is_empty tfrees andalso Names.is_empty frees then th  | 
| 
74200
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1574  | 
else  | 
| 
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1575  | 
let  | 
| 
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1576  | 
      val Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = th;
 | 
| 
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1577  | 
      val _ = idx <= maxidx andalso raise THM ("generalize: bad index", idx, [th]);
 | 
| 19910 | 1578  | 
|
| 
74200
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1579  | 
val bad_type =  | 
| 74266 | 1580  | 
if Names.is_empty tfrees then K false  | 
1581  | 
else Term.exists_subtype (fn TFree (a, _) => Names.defined tfrees a | _ => false);  | 
|
1582  | 
fun bad_term (Free (x, T)) = bad_type T orelse Names.defined frees x  | 
|
| 
74200
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1583  | 
| bad_term (Var (_, T)) = bad_type T  | 
| 
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1584  | 
| bad_term (Const (_, T)) = bad_type T  | 
| 
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1585  | 
| bad_term (Abs (_, T, t)) = bad_type T orelse bad_term t  | 
| 
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1586  | 
| bad_term (t $ u) = bad_term t orelse bad_term u  | 
| 
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1587  | 
| bad_term (Bound _) = false;  | 
| 
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1588  | 
val _ = exists bad_term hyps andalso  | 
| 
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1589  | 
        raise THM ("generalize: variable free in assumptions", 0, [th]);
 | 
| 19910 | 1590  | 
|
| 
74200
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1591  | 
val generalize = Term_Subst.generalize (tfrees, frees) idx;  | 
| 
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1592  | 
val prop' = generalize prop;  | 
| 
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1593  | 
val tpairs' = map (apply2 generalize) tpairs;  | 
| 
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1594  | 
val maxidx' = maxidx_tpairs tpairs' (maxidx_of_term prop');  | 
| 
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1595  | 
in  | 
| 
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1596  | 
Thm (deriv_rule1 (Proofterm.generalize_proof (tfrees, frees) idx prop) der,  | 
| 
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1597  | 
       {cert = cert,
 | 
| 
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1598  | 
tags = [],  | 
| 
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1599  | 
maxidx = maxidx',  | 
| 
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1600  | 
constraints = constraints,  | 
| 
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1601  | 
shyps = shyps,  | 
| 
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1602  | 
hyps = hyps,  | 
| 
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1603  | 
tpairs = tpairs',  | 
| 
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1604  | 
prop = prop'})  | 
| 
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1605  | 
end;  | 
| 19910 | 1606  | 
|
| 
74200
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1607  | 
fun generalize_cterm (tfrees, frees) idx (ct as Cterm {cert, t, T, maxidx, sorts}) =
 | 
| 74266 | 1608  | 
if Names.is_empty tfrees andalso Names.is_empty frees then ct  | 
| 
74200
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1609  | 
  else if idx <= maxidx then raise CTERM ("generalize_cterm: bad index", [ct])
 | 
| 
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1610  | 
else  | 
| 
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1611  | 
    Cterm {cert = cert, sorts = sorts,
 | 
| 
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1612  | 
T = Term_Subst.generalizeT tfrees idx T,  | 
| 
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1613  | 
t = Term_Subst.generalize (tfrees, frees) idx t,  | 
| 
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1614  | 
maxidx = Int.max (maxidx, idx)};  | 
| 67697 | 1615  | 
|
| 
74200
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1616  | 
fun generalize_ctyp tfrees idx (cT as Ctyp {cert, T, maxidx, sorts}) =
 | 
| 74266 | 1617  | 
if Names.is_empty tfrees then cT  | 
| 
74200
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1618  | 
  else if idx <= maxidx then raise CTERM ("generalize_ctyp: bad index", [])
 | 
| 
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1619  | 
else  | 
| 
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1620  | 
    Ctyp {cert = cert, sorts = sorts,
 | 
| 
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1621  | 
T = Term_Subst.generalizeT tfrees idx T,  | 
| 
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
74112 
diff
changeset
 | 
1622  | 
maxidx = Int.max (maxidx, idx)};  | 
| 67697 | 1623  | 
|
| 19910 | 1624  | 
|
| 22584 | 1625  | 
(*Instantiation of schematic variables  | 
| 16656 | 1626  | 
A  | 
1627  | 
--------------------  | 
|
1628  | 
A[t1/v1, ..., tn/vn]  | 
|
| 1220 | 1629  | 
*)  | 
| 0 | 1630  | 
|
| 6928 | 1631  | 
local  | 
1632  | 
||
| 74219 | 1633  | 
fun add_cert cert_of (_, c) cert = Context.join_certificate (cert, cert_of c);  | 
1634  | 
val add_instT_cert = add_cert (fn Ctyp {cert, ...} => cert);
 | 
|
1635  | 
val add_inst_cert = add_cert (fn Cterm {cert, ...} => cert);
 | 
|
1636  | 
||
1637  | 
fun add_sorts sorts_of (_, c) sorts = Sorts.union (sorts_of c) sorts;  | 
|
1638  | 
val add_instT_sorts = add_sorts (fn Ctyp {sorts, ...} => sorts);
 | 
|
1639  | 
val add_inst_sorts = add_sorts (fn Cterm {sorts, ...} => sorts);
 | 
|
1640  | 
||
| 74282 | 1641  | 
fun make_instT thy (v as (_, S)) (Ctyp {T = U, maxidx, ...}) =
 | 
1642  | 
if Sign.of_sort thy (U, S) then (U, maxidx)  | 
|
| 74219 | 1643  | 
  else raise TYPE ("Type not of sort " ^ Syntax.string_of_sort_global thy S, [U], []);
 | 
| 15797 | 1644  | 
|
| 74282 | 1645  | 
fun make_inst thy (v as (_, T)) (Cterm {t = u, T = U, maxidx, ...}) =
 | 
1646  | 
if T = U then (u, maxidx)  | 
|
| 74219 | 1647  | 
else  | 
1648  | 
let  | 
|
1649  | 
fun pretty_typing t ty =  | 
|
1650  | 
Pretty.block [Syntax.pretty_term_global thy t, Pretty.str " ::",  | 
|
1651  | 
Pretty.brk 1, Syntax.pretty_typ_global thy ty];  | 
|
1652  | 
val msg =  | 
|
1653  | 
Pretty.string_of (Pretty.block  | 
|
1654  | 
[Pretty.str "instantiate: type conflict",  | 
|
1655  | 
Pretty.fbrk, pretty_typing (Var v) T,  | 
|
1656  | 
Pretty.fbrk, pretty_typing u U])  | 
|
1657  | 
in raise TYPE (msg, [T, U], [Var v, u]) end;  | 
|
1658  | 
||
1659  | 
fun prep_insts (instT, inst) (cert, sorts) =  | 
|
| 6928 | 1660  | 
let  | 
| 74219 | 1661  | 
val cert' = cert  | 
| 74282 | 1662  | 
|> TVars.fold add_instT_cert instT  | 
1663  | 
|> Vars.fold add_inst_cert inst;  | 
|
1664  | 
val thy' =  | 
|
1665  | 
Context.certificate_theory cert' handle ERROR msg =>  | 
|
1666  | 
raise CONTEXT (msg, TVars.dest instT |> map #2, Vars.dest inst |> map #2, [], NONE);  | 
|
| 0 | 1667  | 
|
| 74219 | 1668  | 
val sorts' = sorts  | 
| 74282 | 1669  | 
|> TVars.fold add_instT_sorts instT  | 
1670  | 
|> Vars.fold add_inst_sorts inst;  | 
|
| 74219 | 1671  | 
|
| 74282 | 1672  | 
val instT' = TVars.map (make_instT thy') instT;  | 
1673  | 
val inst' = Vars.map (make_inst thy') inst;  | 
|
| 74219 | 1674  | 
in ((instT', inst'), (cert', sorts')) end;  | 
| 0 | 1675  | 
|
| 6928 | 1676  | 
in  | 
1677  | 
||
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1678  | 
(*Left-to-right replacements: ctpairs = [..., (vi, ti), ...].  | 
| 0 | 1679  | 
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
 | 
1680  | 
Does NOT normalize the resulting theorem!*)  | 
| 
74577
 
d4829a7333e2
clarified instantiation: local beta reduction after substitution, as for Envir.expand_term_defs;
 
wenzelm 
parents: 
74561 
diff
changeset
 | 
1681  | 
fun gen_instantiate inst_fn (instT, inst) th =  | 
| 74282 | 1682  | 
if TVars.is_empty instT andalso Vars.is_empty inst then th  | 
1683  | 
else  | 
|
1684  | 
let  | 
|
1685  | 
      val Thm (der, {cert, hyps, constraints, shyps, tpairs, prop, ...}) = th;
 | 
|
1686  | 
val ((instT', inst'), (cert', shyps')) = prep_insts (instT, inst) (cert, shyps)  | 
|
1687  | 
handle CONTEXT (msg, cTs, cts, ths, context) =>  | 
|
1688  | 
raise CONTEXT (msg, cTs, cts, th :: ths, context);  | 
|
| 74219 | 1689  | 
|
| 
74577
 
d4829a7333e2
clarified instantiation: local beta reduction after substitution, as for Envir.expand_term_defs;
 
wenzelm 
parents: 
74561 
diff
changeset
 | 
1690  | 
val subst = inst_fn (instT', inst');  | 
| 74282 | 1691  | 
val (prop', maxidx1) = subst prop ~1;  | 
1692  | 
val (tpairs', maxidx') =  | 
|
1693  | 
fold_map (fn (t, u) => fn i => subst t i ||>> subst u) tpairs maxidx1;  | 
|
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1694  | 
|
| 74282 | 1695  | 
val thy' = Context.certificate_theory cert';  | 
1696  | 
val constraints' =  | 
|
1697  | 
TVars.fold (fn ((_, S), (T, _)) => insert_constraints thy' (T, S))  | 
|
1698  | 
instT' constraints;  | 
|
1699  | 
in  | 
|
1700  | 
Thm (deriv_rule1  | 
|
1701  | 
(fn d => Proofterm.instantiate (TVars.map (K #1) instT', Vars.map (K #1) inst') d) der,  | 
|
1702  | 
       {cert = cert',
 | 
|
1703  | 
tags = [],  | 
|
1704  | 
maxidx = maxidx',  | 
|
1705  | 
constraints = constraints',  | 
|
1706  | 
shyps = shyps',  | 
|
1707  | 
hyps = hyps,  | 
|
1708  | 
tpairs = tpairs',  | 
|
1709  | 
prop = prop'})  | 
|
1710  | 
|> solve_constraints  | 
|
1711  | 
end  | 
|
1712  | 
handle TYPE (msg, _, _) => raise THM (msg, 0, [th]);  | 
|
| 6928 | 1713  | 
|
| 
74577
 
d4829a7333e2
clarified instantiation: local beta reduction after substitution, as for Envir.expand_term_defs;
 
wenzelm 
parents: 
74561 
diff
changeset
 | 
1714  | 
val instantiate = gen_instantiate Term_Subst.instantiate_maxidx;  | 
| 
 
d4829a7333e2
clarified instantiation: local beta reduction after substitution, as for Envir.expand_term_defs;
 
wenzelm 
parents: 
74561 
diff
changeset
 | 
1715  | 
val instantiate_beta = gen_instantiate Term_Subst.instantiate_beta_maxidx;  | 
| 
 
d4829a7333e2
clarified instantiation: local beta reduction after substitution, as for Envir.expand_term_defs;
 
wenzelm 
parents: 
74561 
diff
changeset
 | 
1716  | 
|
| 
 
d4829a7333e2
clarified instantiation: local beta reduction after substitution, as for Envir.expand_term_defs;
 
wenzelm 
parents: 
74561 
diff
changeset
 | 
1717  | 
fun gen_instantiate_cterm inst_fn (instT, inst) ct =  | 
| 74282 | 1718  | 
if TVars.is_empty instT andalso Vars.is_empty inst then ct  | 
1719  | 
else  | 
|
1720  | 
let  | 
|
1721  | 
      val Cterm {cert, t, T, sorts, ...} = ct;
 | 
|
1722  | 
val ((instT', inst'), (cert', sorts')) = prep_insts (instT, inst) (cert, sorts);  | 
|
| 
74577
 
d4829a7333e2
clarified instantiation: local beta reduction after substitution, as for Envir.expand_term_defs;
 
wenzelm 
parents: 
74561 
diff
changeset
 | 
1723  | 
val subst = inst_fn (instT', inst');  | 
| 74282 | 1724  | 
val substT = Term_Subst.instantiateT_maxidx instT';  | 
1725  | 
val (t', maxidx1) = subst t ~1;  | 
|
1726  | 
val (T', maxidx') = substT T maxidx1;  | 
|
1727  | 
    in Cterm {cert = cert', t = t', T = T', sorts = sorts', maxidx = maxidx'} end
 | 
|
1728  | 
handle TYPE (msg, _, _) => raise CTERM (msg, [ct]);  | 
|
| 22584 | 1729  | 
|
| 
74577
 
d4829a7333e2
clarified instantiation: local beta reduction after substitution, as for Envir.expand_term_defs;
 
wenzelm 
parents: 
74561 
diff
changeset
 | 
1730  | 
val instantiate_cterm = gen_instantiate_cterm Term_Subst.instantiate_maxidx;  | 
| 
 
d4829a7333e2
clarified instantiation: local beta reduction after substitution, as for Envir.expand_term_defs;
 
wenzelm 
parents: 
74561 
diff
changeset
 | 
1731  | 
val instantiate_beta_cterm = gen_instantiate_cterm Term_Subst.instantiate_beta_maxidx;  | 
| 
 
d4829a7333e2
clarified instantiation: local beta reduction after substitution, as for Envir.expand_term_defs;
 
wenzelm 
parents: 
74561 
diff
changeset
 | 
1732  | 
|
| 6928 | 1733  | 
end;  | 
1734  | 
||
| 0 | 1735  | 
|
| 67721 | 1736  | 
(*The trivial implication A \<Longrightarrow> A, justified by assume and forall rules.  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1737  | 
A can contain Vars, not so for assume!*)  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
1738  | 
fun trivial (Cterm {cert, t = A, T, maxidx, sorts}) =
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1739  | 
if T <> propT then  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1740  | 
    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
 | 
1741  | 
else  | 
| 70824 | 1742  | 
Thm (deriv_rule0 (fn () => Proofterm.trivial_proof),  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
1743  | 
     {cert = cert,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
1744  | 
tags = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1745  | 
maxidx = maxidx,  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1746  | 
constraints = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1747  | 
shyps = sorts,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1748  | 
hyps = [],  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1749  | 
tpairs = [],  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1750  | 
prop = Logic.mk_implies (A, A)});  | 
| 0 | 1751  | 
|
| 31944 | 1752  | 
(*Axiom-scheme reflecting signature contents  | 
1753  | 
T :: c  | 
|
1754  | 
-------------------  | 
|
1755  | 
OFCLASS(T, c_class)  | 
|
1756  | 
*)  | 
|
1757  | 
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
 | 
1758  | 
let  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
1759  | 
    val Ctyp {cert, T, ...} = cT;
 | 
| 61051 | 1760  | 
val thy = Context.certificate_theory cert  | 
1761  | 
handle ERROR msg => raise CONTEXT (msg, [cT], [], [], NONE);  | 
|
| 31903 | 1762  | 
val c = Sign.certify_class thy raw_c;  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59591 
diff
changeset
 | 
1763  | 
    val Cterm {t = prop, maxidx, sorts, ...} = global_cterm_of thy (Logic.mk_of_class (T, c));
 | 
| 
399
 
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
 
wenzelm 
parents: 
387 
diff
changeset
 | 
1764  | 
in  | 
| 31944 | 1765  | 
if Sign.of_sort thy (T, [c]) then  | 
| 
71777
 
3875815f5967
clarified signature: avoid clash with Isabelle/Scala Term.OFCLASS on case-insensible file-system;
 
wenzelm 
parents: 
71553 
diff
changeset
 | 
1766  | 
Thm (deriv_rule0 (fn () => Proofterm.PClass (T, c)),  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
1767  | 
       {cert = cert,
 | 
| 31944 | 1768  | 
tags = [],  | 
1769  | 
maxidx = maxidx,  | 
|
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1770  | 
constraints = insert_constraints thy (T, [c]) [],  | 
| 31944 | 1771  | 
shyps = sorts,  | 
1772  | 
hyps = [],  | 
|
1773  | 
tpairs = [],  | 
|
1774  | 
prop = prop})  | 
|
1775  | 
    else raise THM ("of_class: type not of class " ^ Syntax.string_of_sort_global thy [c], 0, [])
 | 
|
| 
70482
 
d4b5139eea34
more robust and convenient treatment of implicit context;
 
wenzelm 
parents: 
70480 
diff
changeset
 | 
1776  | 
end |> solve_constraints;  | 
| 
399
 
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
 
wenzelm 
parents: 
387 
diff
changeset
 | 
1777  | 
|
| 
36769
 
b6b88bf695b3
Thm.unconstrainT based on Logic.unconstrainT -- no proofterm yet;
 
wenzelm 
parents: 
36768 
diff
changeset
 | 
1778  | 
(*Internalize sort constraints of type variables*)  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1779  | 
val unconstrainT =  | 
| 
71530
 
046cf49162a3
more thorough strip_shyps for proof boxes (but types are usually stripped and reconstructed later);
 
wenzelm 
parents: 
71527 
diff
changeset
 | 
1780  | 
strip_shyps #> (fn thm as Thm (der, args) =>  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1781  | 
let  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1782  | 
      val Deriv {promises, body} = der;
 | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1783  | 
      val {cert, shyps, hyps, tpairs, prop, ...} = args;
 | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1784  | 
val thy = theory_of_thm thm;  | 
| 
36883
 
4ed0d72def50
added Proofterm.unconstrain_thm_proof for Thm.unconstrainT -- loosely based on the version by krauss/schropp;
 
wenzelm 
parents: 
36882 
diff
changeset
 | 
1785  | 
|
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1786  | 
      fun err msg = raise THM ("unconstrainT: " ^ msg, 0, [thm]);
 | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1787  | 
val _ = null hyps orelse err "bad hyps";  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1788  | 
val _ = null tpairs orelse err "bad flex-flex constraints";  | 
| 74235 | 1789  | 
val tfrees = build_rev (Term.add_tfree_names prop);  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1790  | 
      val _ = null tfrees orelse err ("illegal free type variables " ^ commas_quote tfrees);
 | 
| 
36769
 
b6b88bf695b3
Thm.unconstrainT based on Logic.unconstrainT -- no proofterm yet;
 
wenzelm 
parents: 
36768 
diff
changeset
 | 
1791  | 
|
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1792  | 
val ps = map (apsnd (Future.map fulfill_body)) promises;  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1793  | 
val (pthm, proof) =  | 
| 70494 | 1794  | 
Proofterm.unconstrain_thm_proof thy (classrel_proof thy) (arity_proof thy)  | 
1795  | 
shyps prop ps body;  | 
|
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1796  | 
val der' = make_deriv [] [] [pthm] proof;  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1797  | 
val prop' = Proofterm.thm_node_prop (#2 pthm);  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1798  | 
in  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1799  | 
Thm (der',  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1800  | 
       {cert = cert,
 | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1801  | 
tags = [],  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1802  | 
maxidx = maxidx_of_term prop',  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1803  | 
constraints = [],  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1804  | 
shyps = [[]], (*potentially redundant*)  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1805  | 
hyps = [],  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1806  | 
tpairs = [],  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1807  | 
prop = prop'})  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1808  | 
end);  | 
| 
399
 
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
 
wenzelm 
parents: 
387 
diff
changeset
 | 
1809  | 
|
| 63611 | 1810  | 
(*Replace all TFrees not fixed or in the hyps by new TVars*)  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1811  | 
fun varifyT_global' fixed (Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...})) =
 | 
| 12500 | 1812  | 
let  | 
| 74266 | 1813  | 
val tfrees = fold TFrees.add_tfrees hyps fixed;  | 
| 
13658
 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 
berghofe 
parents: 
13642 
diff
changeset
 | 
1814  | 
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
 | 
1815  | 
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
 | 
1816  | 
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
 | 
1817  | 
in  | 
| 
52487
 
48bc24467008
backout dedd7952a62c: static "proofs" value within theory prevents later inferencing with different configuration;
 
wenzelm 
parents: 
52470 
diff
changeset
 | 
1818  | 
(al, Thm (deriv_rule1 (Proofterm.varify_proof prop tfrees) der,  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
1819  | 
     {cert = cert,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
1820  | 
tags = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1821  | 
maxidx = Int.max (0, maxidx),  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1822  | 
constraints = constraints,  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1823  | 
shyps = shyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1824  | 
hyps = hyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1825  | 
tpairs = rev (map Logic.dest_equals ts),  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1826  | 
prop = prop3}))  | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1827  | 
end;  | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1828  | 
|
| 74266 | 1829  | 
val varifyT_global = #2 o varifyT_global' TFrees.empty;  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1830  | 
|
| 63611 | 1831  | 
(*Replace all TVars by TFrees that are often new*)  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1832  | 
fun legacy_freezeT (Thm (der, {cert, constraints, shyps, hyps, tpairs, prop, ...})) =
 | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1833  | 
let  | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1834  | 
val prop1 = attach_tpairs tpairs prop;  | 
| 33832 | 1835  | 
val prop2 = Type.legacy_freeze prop1;  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1836  | 
val (ts, prop3) = Logic.strip_prems (length tpairs, [], prop2);  | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1837  | 
in  | 
| 
52487
 
48bc24467008
backout dedd7952a62c: static "proofs" value within theory prevents later inferencing with different configuration;
 
wenzelm 
parents: 
52470 
diff
changeset
 | 
1838  | 
Thm (deriv_rule1 (Proofterm.legacy_freezeT prop1) der,  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
1839  | 
     {cert = cert,
 | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1840  | 
tags = [],  | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1841  | 
maxidx = maxidx_of_term prop2,  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1842  | 
constraints = constraints,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1843  | 
shyps = shyps,  | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1844  | 
hyps = hyps,  | 
| 
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1845  | 
tpairs = rev (map Logic.dest_equals ts),  | 
| 18127 | 1846  | 
prop = prop3})  | 
| 0 | 1847  | 
end;  | 
1848  | 
||
| 
70456
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
1849  | 
fun plain_prop_of raw_thm =  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
1850  | 
let  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
1851  | 
val thm = strip_shyps raw_thm;  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
1852  | 
    fun err msg = raise THM ("plain_prop_of: " ^ msg, 0, [thm]);
 | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
1853  | 
in  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
1854  | 
if not (null (hyps_of thm)) then  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
1855  | 
err "theorem may not contain hypotheses"  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
1856  | 
else if not (null (extra_shyps thm)) then  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
1857  | 
err "theorem may not contain sort hypotheses"  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
1858  | 
else if not (null (tpairs_of thm)) then  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
1859  | 
err "theorem may not contain flex-flex pairs"  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
1860  | 
else prop_of thm  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
1861  | 
end;  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
1862  | 
|
| 0 | 1863  | 
|
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1864  | 
|
| 0 | 1865  | 
(*** Inference rules for tactics ***)  | 
1866  | 
||
1867  | 
(*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
 | 
1868  | 
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
 | 
1869  | 
(case Logic.strip_prems(i, [], prop) of  | 
| 
 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 
berghofe 
parents: 
13642 
diff
changeset
 | 
1870  | 
(B::rBs, C) => (tpairs, rev rBs, B, C)  | 
| 
 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 
berghofe 
parents: 
13642 
diff
changeset
 | 
1871  | 
    | _ => raise THM("dest_state", i, [state]))
 | 
| 0 | 1872  | 
  handle TERM _ => raise THM("dest_state", i, [state]);
 | 
1873  | 
||
| 46255 | 1874  | 
(*Prepare orule for resolution by lifting it over the parameters and  | 
| 
52487
 
48bc24467008
backout dedd7952a62c: static "proofs" value within theory prevents later inferencing with different configuration;
 
wenzelm 
parents: 
52470 
diff
changeset
 | 
1875  | 
assumptions of goal.*)  | 
| 18035 | 1876  | 
fun lift_rule goal orule =  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1877  | 
let  | 
| 18035 | 1878  | 
    val Cterm {t = gprop, T, maxidx = gmax, sorts, ...} = goal;
 | 
1879  | 
val inc = gmax + 1;  | 
|
1880  | 
val lift_abs = Logic.lift_abs inc gprop;  | 
|
1881  | 
val lift_all = Logic.lift_all inc gprop;  | 
|
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1882  | 
    val Thm (der, {maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = orule;
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1883  | 
val (As, B) = Logic.strip_horn prop;  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1884  | 
in  | 
| 18035 | 1885  | 
    if T <> propT then raise THM ("lift_rule: the term must have type prop", 0, [])
 | 
1886  | 
else  | 
|
| 
52487
 
48bc24467008
backout dedd7952a62c: static "proofs" value within theory prevents later inferencing with different configuration;
 
wenzelm 
parents: 
52470 
diff
changeset
 | 
1887  | 
Thm (deriv_rule1 (Proofterm.lift_proof gprop inc prop) der,  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
1888  | 
       {cert = join_certificate1 (goal, orule),
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
1889  | 
tags = [],  | 
| 18035 | 1890  | 
maxidx = maxidx + inc,  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1891  | 
constraints = constraints,  | 
| 18035 | 1892  | 
shyps = Sorts.union shyps sorts, (*sic!*)  | 
1893  | 
hyps = hyps,  | 
|
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58950 
diff
changeset
 | 
1894  | 
tpairs = map (apply2 lift_abs) tpairs,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1895  | 
prop = Logic.list_implies (map lift_all As, lift_all B)})  | 
| 0 | 1896  | 
end;  | 
1897  | 
||
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1898  | 
fun incr_indexes i (thm as Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...})) =
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1899  | 
  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
 | 
1900  | 
else if i = 0 then thm  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1901  | 
else  | 
| 
52487
 
48bc24467008
backout dedd7952a62c: static "proofs" value within theory prevents later inferencing with different configuration;
 
wenzelm 
parents: 
52470 
diff
changeset
 | 
1902  | 
Thm (deriv_rule1 (Proofterm.incr_indexes i) der,  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
1903  | 
     {cert = cert,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
1904  | 
tags = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1905  | 
maxidx = maxidx + i,  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1906  | 
constraints = constraints,  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1907  | 
shyps = shyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1908  | 
hyps = hyps,  | 
| 59787 | 1909  | 
tpairs = map (apply2 (Logic.incr_indexes ([], [], i))) tpairs,  | 
1910  | 
prop = Logic.incr_indexes ([], [], i) prop});  | 
|
| 
10416
 
5b33e732e459
- Moved rewriting functions to meta_simplifier.ML
 
berghofe 
parents: 
10348 
diff
changeset
 | 
1911  | 
|
| 0 | 1912  | 
(*Solve subgoal Bi of proof state B1...Bn/C by assumption. *)  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
58946 
diff
changeset
 | 
1913  | 
fun assumption opt_ctxt i state =  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1914  | 
let  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1915  | 
    val Thm (der, {cert, maxidx, constraints, shyps, hyps, ...}) = state;
 | 
| 
64981
 
ea6199b23dfa
proper background certificate from make_context, which can be a super-theory of the direct join (amending d07464875dd4);
 
wenzelm 
parents: 
64574 
diff
changeset
 | 
1916  | 
val (context, cert') = make_context_certificate [state] opt_ctxt cert;  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1917  | 
val (tpairs, Bs, Bi, C) = dest_state (state, i);  | 
| 32032 | 1918  | 
fun newth n (env, tpairs) =  | 
| 
70822
 
22e2f5acc0b4
more accurate treatment of propositions within proof terms, but these are ultimately ignored for performance reasons;
 
wenzelm 
parents: 
70814 
diff
changeset
 | 
1919  | 
let  | 
| 
 
22e2f5acc0b4
more accurate treatment of propositions within proof terms, but these are ultimately ignored for performance reasons;
 
wenzelm 
parents: 
70814 
diff
changeset
 | 
1920  | 
val normt = Envir.norm_term env;  | 
| 
 
22e2f5acc0b4
more accurate treatment of propositions within proof terms, but these are ultimately ignored for performance reasons;
 
wenzelm 
parents: 
70814 
diff
changeset
 | 
1921  | 
fun assumption_proof prf =  | 
| 70825 | 1922  | 
Proofterm.assumption_proof (map normt Bs) (normt Bi) n prf;  | 
| 
70822
 
22e2f5acc0b4
more accurate treatment of propositions within proof terms, but these are ultimately ignored for performance reasons;
 
wenzelm 
parents: 
70814 
diff
changeset
 | 
1923  | 
in  | 
| 70825 | 1924  | 
Thm (deriv_rule1  | 
1925  | 
(assumption_proof #> not (Envir.is_empty env) ? Proofterm.norm_proof' env) der,  | 
|
| 
70822
 
22e2f5acc0b4
more accurate treatment of propositions within proof terms, but these are ultimately ignored for performance reasons;
 
wenzelm 
parents: 
70814 
diff
changeset
 | 
1926  | 
         {tags = [],
 | 
| 
 
22e2f5acc0b4
more accurate treatment of propositions within proof terms, but these are ultimately ignored for performance reasons;
 
wenzelm 
parents: 
70814 
diff
changeset
 | 
1927  | 
maxidx = Envir.maxidx_of env,  | 
| 
 
22e2f5acc0b4
more accurate treatment of propositions within proof terms, but these are ultimately ignored for performance reasons;
 
wenzelm 
parents: 
70814 
diff
changeset
 | 
1928  | 
constraints = insert_constraints_env (Context.certificate_theory cert') env constraints,  | 
| 
 
22e2f5acc0b4
more accurate treatment of propositions within proof terms, but these are ultimately ignored for performance reasons;
 
wenzelm 
parents: 
70814 
diff
changeset
 | 
1929  | 
shyps = Envir.insert_sorts env shyps,  | 
| 
 
22e2f5acc0b4
more accurate treatment of propositions within proof terms, but these are ultimately ignored for performance reasons;
 
wenzelm 
parents: 
70814 
diff
changeset
 | 
1930  | 
hyps = hyps,  | 
| 
 
22e2f5acc0b4
more accurate treatment of propositions within proof terms, but these are ultimately ignored for performance reasons;
 
wenzelm 
parents: 
70814 
diff
changeset
 | 
1931  | 
tpairs = if Envir.is_empty env then tpairs else map (apply2 normt) tpairs,  | 
| 
 
22e2f5acc0b4
more accurate treatment of propositions within proof terms, but these are ultimately ignored for performance reasons;
 
wenzelm 
parents: 
70814 
diff
changeset
 | 
1932  | 
prop =  | 
| 
 
22e2f5acc0b4
more accurate treatment of propositions within proof terms, but these are ultimately ignored for performance reasons;
 
wenzelm 
parents: 
70814 
diff
changeset
 | 
1933  | 
if Envir.is_empty env then Logic.list_implies (Bs, C) (*avoid wasted normalizations*)  | 
| 
 
22e2f5acc0b4
more accurate treatment of propositions within proof terms, but these are ultimately ignored for performance reasons;
 
wenzelm 
parents: 
70814 
diff
changeset
 | 
1934  | 
else normt (Logic.list_implies (Bs, C)) (*normalize the new rule fully*),  | 
| 
 
22e2f5acc0b4
more accurate treatment of propositions within proof terms, but these are ultimately ignored for performance reasons;
 
wenzelm 
parents: 
70814 
diff
changeset
 | 
1935  | 
cert = cert'})  | 
| 
 
22e2f5acc0b4
more accurate treatment of propositions within proof terms, but these are ultimately ignored for performance reasons;
 
wenzelm 
parents: 
70814 
diff
changeset
 | 
1936  | 
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
 | 
1937  | 
|
| 
30556
 
7be15917f3fa
eq_assumption: slightly more efficient by checking (open) result of Logic.assum_problems directly;
 
wenzelm 
parents: 
30554 
diff
changeset
 | 
1938  | 
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
 | 
1939  | 
val concl' = close concl;  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1940  | 
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
 | 
1941  | 
| 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
 | 
1942  | 
(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
 | 
1943  | 
(if Term.could_unify (asm, concl) then  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
58946 
diff
changeset
 | 
1944  | 
(Unify.unifiers (context, 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
 | 
1945  | 
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
 | 
1946  | 
(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
 | 
1947  | 
in addprfs asms 1 end;  | 
| 0 | 1948  | 
|
| 250 | 1949  | 
(*Solve subgoal Bi of proof state B1...Bn/C by assumption.  | 
| 51604 | 1950  | 
Checks if Bi's conclusion is alpha/eta-convertible to one of its assumptions*)  | 
| 0 | 1951  | 
fun eq_assumption i state =  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1952  | 
let  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1953  | 
    val Thm (der, {cert, maxidx, constraints, shyps, hyps, ...}) = state;
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1954  | 
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
 | 
1955  | 
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
 | 
1956  | 
in  | 
| 52131 | 1957  | 
(case find_index (fn asm => Envir.aeconv (asm, concl)) asms of  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1958  | 
      ~1 => raise THM ("eq_assumption", 0, [state])
 | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1959  | 
| n =>  | 
| 
52487
 
48bc24467008
backout dedd7952a62c: static "proofs" value within theory prevents later inferencing with different configuration;
 
wenzelm 
parents: 
52470 
diff
changeset
 | 
1960  | 
Thm (deriv_rule1 (Proofterm.assumption_proof Bs Bi (n + 1)) der,  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
1961  | 
         {cert = cert,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
1962  | 
tags = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1963  | 
maxidx = maxidx,  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1964  | 
constraints = constraints,  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1965  | 
shyps = shyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1966  | 
hyps = hyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1967  | 
tpairs = tpairs,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1968  | 
prop = Logic.list_implies (Bs, C)}))  | 
| 0 | 1969  | 
end;  | 
1970  | 
||
1971  | 
||
| 2671 | 1972  | 
(*For rotate_tac: fast rotation of assumptions of subgoal i*)  | 
1973  | 
fun rotate_rule k i state =  | 
|
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1974  | 
let  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1975  | 
    val Thm (der, {cert, maxidx, constraints, shyps, hyps, ...}) = state;
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1976  | 
val (tpairs, Bs, Bi, C) = dest_state (state, i);  | 
| 
46218
 
ecf6375e2abb
renamed Term.list_all to Logic.list_all, in accordance to HOLogic.list_all;
 
wenzelm 
parents: 
46217 
diff
changeset
 | 
1977  | 
val params = Term.strip_all_vars Bi;  | 
| 
 
ecf6375e2abb
renamed Term.list_all to Logic.list_all, in accordance to HOLogic.list_all;
 
wenzelm 
parents: 
46217 
diff
changeset
 | 
1978  | 
val rest = Term.strip_all_body Bi;  | 
| 
 
ecf6375e2abb
renamed Term.list_all to Logic.list_all, in accordance to HOLogic.list_all;
 
wenzelm 
parents: 
46217 
diff
changeset
 | 
1979  | 
val asms = Logic.strip_imp_prems rest  | 
| 
 
ecf6375e2abb
renamed Term.list_all to Logic.list_all, in accordance to HOLogic.list_all;
 
wenzelm 
parents: 
46217 
diff
changeset
 | 
1980  | 
val concl = Logic.strip_imp_concl rest;  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1981  | 
val n = length asms;  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1982  | 
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
 | 
1983  | 
val Bi' =  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1984  | 
if 0 = m orelse m = n then Bi  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1985  | 
else if 0 < m andalso m < n then  | 
| 19012 | 1986  | 
let val (ps, qs) = chop m asms  | 
| 
46218
 
ecf6375e2abb
renamed Term.list_all to Logic.list_all, in accordance to HOLogic.list_all;
 
wenzelm 
parents: 
46217 
diff
changeset
 | 
1987  | 
in Logic.list_all (params, Logic.list_implies (qs @ ps, concl)) end  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1988  | 
      else raise THM ("rotate_rule", k, [state]);
 | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1989  | 
in  | 
| 
70822
 
22e2f5acc0b4
more accurate treatment of propositions within proof terms, but these are ultimately ignored for performance reasons;
 
wenzelm 
parents: 
70814 
diff
changeset
 | 
1990  | 
Thm (deriv_rule1 (Proofterm.rotate_proof Bs Bi' params asms m) der,  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
1991  | 
     {cert = cert,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
1992  | 
tags = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1993  | 
maxidx = maxidx,  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
1994  | 
constraints = constraints,  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1995  | 
shyps = shyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1996  | 
hyps = hyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
1997  | 
tpairs = tpairs,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
1998  | 
prop = Logic.list_implies (Bs @ [Bi'], C)})  | 
| 2671 | 1999  | 
end;  | 
2000  | 
||
2001  | 
||
| 
7248
 
322151fe6f02
new primitive rule permute_prems to underlie defer_tac and rotate_prems
 
paulson 
parents: 
7070 
diff
changeset
 | 
2002  | 
(*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
 | 
2003  | 
unchanged. Does nothing if k=0 or if k equals n-j, where n is the  | 
| 58837 | 2004  | 
number of premises. Useful with eresolve_tac and underlies defer_tac*)  | 
| 
7248
 
322151fe6f02
new primitive rule permute_prems to underlie defer_tac and rotate_prems
 
paulson 
parents: 
7070 
diff
changeset
 | 
2005  | 
fun permute_prems j k rl =  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
2006  | 
let  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
2007  | 
    val Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = rl;
 | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
2008  | 
val prems = Logic.strip_imp_prems prop  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
2009  | 
and concl = Logic.strip_imp_concl prop;  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
2010  | 
val moved_prems = List.drop (prems, j)  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
2011  | 
and fixed_prems = List.take (prems, j)  | 
| 43278 | 2012  | 
      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
 | 
2013  | 
val n_j = length moved_prems;  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
2014  | 
val m = if k < 0 then n_j + k else k;  | 
| 
70822
 
22e2f5acc0b4
more accurate treatment of propositions within proof terms, but these are ultimately ignored for performance reasons;
 
wenzelm 
parents: 
70814 
diff
changeset
 | 
2015  | 
val (prems', prop') =  | 
| 
 
22e2f5acc0b4
more accurate treatment of propositions within proof terms, but these are ultimately ignored for performance reasons;
 
wenzelm 
parents: 
70814 
diff
changeset
 | 
2016  | 
if 0 = m orelse m = n_j then (prems, prop)  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
2017  | 
else if 0 < m andalso m < n_j then  | 
| 
70822
 
22e2f5acc0b4
more accurate treatment of propositions within proof terms, but these are ultimately ignored for performance reasons;
 
wenzelm 
parents: 
70814 
diff
changeset
 | 
2018  | 
let  | 
| 
 
22e2f5acc0b4
more accurate treatment of propositions within proof terms, but these are ultimately ignored for performance reasons;
 
wenzelm 
parents: 
70814 
diff
changeset
 | 
2019  | 
val (ps, qs) = chop m moved_prems;  | 
| 
 
22e2f5acc0b4
more accurate treatment of propositions within proof terms, but these are ultimately ignored for performance reasons;
 
wenzelm 
parents: 
70814 
diff
changeset
 | 
2020  | 
val prems' = fixed_prems @ qs @ ps;  | 
| 
 
22e2f5acc0b4
more accurate treatment of propositions within proof terms, but these are ultimately ignored for performance reasons;
 
wenzelm 
parents: 
70814 
diff
changeset
 | 
2021  | 
in (prems', Logic.list_implies (prems', concl)) end  | 
| 16725 | 2022  | 
      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
 | 
2023  | 
in  | 
| 
70822
 
22e2f5acc0b4
more accurate treatment of propositions within proof terms, but these are ultimately ignored for performance reasons;
 
wenzelm 
parents: 
70814 
diff
changeset
 | 
2024  | 
Thm (deriv_rule1 (Proofterm.permute_prems_proof prems' j m) der,  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
2025  | 
     {cert = cert,
 | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
21576 
diff
changeset
 | 
2026  | 
tags = [],  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
2027  | 
maxidx = maxidx,  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
2028  | 
constraints = constraints,  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
2029  | 
shyps = shyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
2030  | 
hyps = hyps,  | 
| 
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
2031  | 
tpairs = tpairs,  | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
2032  | 
prop = prop'})  | 
| 
7248
 
322151fe6f02
new primitive rule permute_prems to underlie defer_tac and rotate_prems
 
paulson 
parents: 
7070 
diff
changeset
 | 
2033  | 
end;  | 
| 
 
322151fe6f02
new primitive rule permute_prems to underlie defer_tac and rotate_prems
 
paulson 
parents: 
7070 
diff
changeset
 | 
2034  | 
|
| 
 
322151fe6f02
new primitive rule permute_prems to underlie defer_tac and rotate_prems
 
paulson 
parents: 
7070 
diff
changeset
 | 
2035  | 
|
| 
44108
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
2036  | 
(* strip_apply f B A strips off all assumptions/parameters from A  | 
| 0 | 2037  | 
introduced by lifting over B, and applies f to remaining part of A*)  | 
2038  | 
fun strip_apply f =  | 
|
| 56245 | 2039  | 
  let fun strip (Const ("Pure.imp", _) $ _  $ B1)
 | 
2040  | 
                (Const ("Pure.imp", _) $ A2 $ B2) = Logic.mk_implies (A2, strip B1 B2)
 | 
|
2041  | 
        | strip ((c as Const ("Pure.all", _)) $ Abs (_, _, t1))
 | 
|
2042  | 
                (      Const ("Pure.all", _)  $ Abs (a, T, t2)) = c $ Abs (a, T, strip t1 t2)
 | 
|
| 
44108
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
2043  | 
| strip _ A = f A  | 
| 0 | 2044  | 
in strip end;  | 
2045  | 
||
| 56245 | 2046  | 
fun strip_lifted (Const ("Pure.imp", _) $ _ $ B1)
 | 
2047  | 
                 (Const ("Pure.imp", _) $ _ $ B2) = strip_lifted B1 B2
 | 
|
2048  | 
  | strip_lifted (Const ("Pure.all", _) $ Abs (_, _, t1))
 | 
|
2049  | 
                 (Const ("Pure.all", _) $ Abs (_, _, t2)) = strip_lifted t1 t2
 | 
|
| 
44108
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
2050  | 
| strip_lifted _ A = A;  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
2051  | 
|
| 0 | 2052  | 
(*Use the alist to rename all bound variables and some unknowns in a term  | 
2053  | 
dpairs = current disagreement pairs; tpairs = permanent ones (flexflex);  | 
|
2054  | 
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
 | 
2055  | 
fun rename_bvs [] _ _ _ _ = K I  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
2056  | 
| rename_bvs al dpairs tpairs B As =  | 
| 20330 | 2057  | 
let  | 
2058  | 
val add_var = fold_aterms (fn Var ((x, _), _) => insert (op =) x | _ => I);  | 
|
2059  | 
val vids = []  | 
|
2060  | 
|> fold (add_var o fst) dpairs  | 
|
2061  | 
|> fold (add_var o fst) tpairs  | 
|
2062  | 
|> fold (add_var o snd) tpairs;  | 
|
| 
44108
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
2063  | 
val vids' = fold (add_var o strip_lifted B) As [];  | 
| 250 | 2064  | 
(*unknowns appearing elsewhere be preserved!*)  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58950 
diff
changeset
 | 
2065  | 
val al' = distinct ((op =) o apply2 fst)  | 
| 
44108
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
2066  | 
(filter_out (fn (x, y) =>  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
2067  | 
not (member (op =) vids' x) orelse  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
2068  | 
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
 | 
2069  | 
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
 | 
2070  | 
fun del_clashing clash xs _ [] qs =  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
2071  | 
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
 | 
2072  | 
| 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
 | 
2073  | 
if member (op =) ys y  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
2074  | 
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
 | 
2075  | 
else del_clashing clash xs (y :: ys) ps (p :: qs);  | 
| 
46497
 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 
wenzelm 
parents: 
46475 
diff
changeset
 | 
2076  | 
val al'' = del_clashing false unchanged unchanged al' [];  | 
| 
44108
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
2077  | 
fun rename (t as Var ((x, i), T)) =  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
2078  | 
(case AList.lookup (op =) al'' x of  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
2079  | 
SOME y => Var ((y, i), T)  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
2080  | 
| NONE => t)  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
2081  | 
| rename (Abs (x, T, t)) =  | 
| 18944 | 2082  | 
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
 | 
2083  | 
| rename (f $ t) = rename f $ rename t  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
2084  | 
| rename t = t;  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
2085  | 
fun strip_ren f Ai = f rename B Ai  | 
| 20330 | 2086  | 
in strip_ren end;  | 
| 0 | 2087  | 
|
2088  | 
(*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
 | 
2089  | 
fun rename_bvars dpairs =  | 
| 48263 | 2090  | 
rename_bvs (fold_rev Term.match_bvars dpairs []) dpairs;  | 
| 0 | 2091  | 
|
2092  | 
||
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
2093  | 
|
| 0 | 2094  | 
(*** RESOLUTION ***)  | 
2095  | 
||
| 
721
 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 
lcp 
parents: 
678 
diff
changeset
 | 
2096  | 
(** Lifting optimizations **)  | 
| 
 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 
lcp 
parents: 
678 
diff
changeset
 | 
2097  | 
|
| 0 | 2098  | 
(*strip off pairs of assumptions/parameters in parallel -- they are  | 
2099  | 
identical because of lifting*)  | 
|
| 56245 | 2100  | 
fun strip_assums2 (Const("Pure.imp", _) $ _ $ B1,
 | 
2101  | 
                   Const("Pure.imp", _) $ _ $ B2) = strip_assums2 (B1,B2)
 | 
|
2102  | 
  | strip_assums2 (Const("Pure.all",_)$Abs(a,T,t1),
 | 
|
2103  | 
                   Const("Pure.all",_)$Abs(_,_,t2)) =
 | 
|
| 0 | 2104  | 
let val (B1,B2) = strip_assums2 (t1,t2)  | 
2105  | 
in (Abs(a,T,B1), Abs(a,T,B2)) end  | 
|
2106  | 
| strip_assums2 BB = BB;  | 
|
2107  | 
||
2108  | 
||
| 
721
 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 
lcp 
parents: 
678 
diff
changeset
 | 
2109  | 
(*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
 | 
2110  | 
fun norm_term_skip env 0 t = Envir.norm_term env t  | 
| 56245 | 2111  | 
  | norm_term_skip env n (Const ("Pure.all", _) $ Abs (a, T, t)) =
 | 
| 32032 | 2112  | 
let  | 
| 
58946
 
3bf80312508e
proper Envir.norm_type for result of Unify.unifiers (amending 479832ff2d29 from 20 years ago);
 
wenzelm 
parents: 
58837 
diff
changeset
 | 
2113  | 
val T' = Envir.norm_type (Envir.type_env env) T  | 
| 32032 | 2114  | 
(*Must instantiate types of parameters because they are flattened;  | 
2115  | 
this could be a NEW parameter*)  | 
|
| 
46217
 
7b19666f0e3d
renamed Term.all to Logic.all_const, in accordance to HOLogic.all_const;
 
wenzelm 
parents: 
45443 
diff
changeset
 | 
2116  | 
in Logic.all_const T' $ Abs (a, T', norm_term_skip env n t) end  | 
| 56245 | 2117  | 
  | norm_term_skip env n (Const ("Pure.imp", _) $ A $ B) =
 | 
| 32032 | 2118  | 
Logic.mk_implies (A, norm_term_skip env (n - 1) B)  | 
| 32784 | 2119  | 
| 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
 | 
2120  | 
|
| 
 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 
lcp 
parents: 
678 
diff
changeset
 | 
2121  | 
|
| 
52222
 
0fa3b456a267
unify types of schematic variables in non-lifted case (i.e. "compose variants") -- allow schematic polymorphism, without revisiting HO-unification;
 
wenzelm 
parents: 
52131 
diff
changeset
 | 
2122  | 
(*unify types of schematic variables (non-lifted case)*)  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
58946 
diff
changeset
 | 
2123  | 
fun unify_var_types context (th1, th2) env =  | 
| 
52222
 
0fa3b456a267
unify types of schematic variables in non-lifted case (i.e. "compose variants") -- allow schematic polymorphism, without revisiting HO-unification;
 
wenzelm 
parents: 
52131 
diff
changeset
 | 
2124  | 
let  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
58946 
diff
changeset
 | 
2125  | 
fun unify_vars (T :: Us) = fold (fn U => Pattern.unify_types context (T, U)) Us  | 
| 
52222
 
0fa3b456a267
unify types of schematic variables in non-lifted case (i.e. "compose variants") -- allow schematic polymorphism, without revisiting HO-unification;
 
wenzelm 
parents: 
52131 
diff
changeset
 | 
2126  | 
| unify_vars _ = I;  | 
| 
 
0fa3b456a267
unify types of schematic variables in non-lifted case (i.e. "compose variants") -- allow schematic polymorphism, without revisiting HO-unification;
 
wenzelm 
parents: 
52131 
diff
changeset
 | 
2127  | 
val add_vars =  | 
| 
 
0fa3b456a267
unify types of schematic variables in non-lifted case (i.e. "compose variants") -- allow schematic polymorphism, without revisiting HO-unification;
 
wenzelm 
parents: 
52131 
diff
changeset
 | 
2128  | 
full_prop_of #>  | 
| 
 
0fa3b456a267
unify types of schematic variables in non-lifted case (i.e. "compose variants") -- allow schematic polymorphism, without revisiting HO-unification;
 
wenzelm 
parents: 
52131 
diff
changeset
 | 
2129  | 
fold_aterms (fn Var v => Vartab.insert_list (op =) v | _ => I);  | 
| 74232 | 2130  | 
val vars = Vartab.build (add_vars th1 #> add_vars th2);  | 
| 
52222
 
0fa3b456a267
unify types of schematic variables in non-lifted case (i.e. "compose variants") -- allow schematic polymorphism, without revisiting HO-unification;
 
wenzelm 
parents: 
52131 
diff
changeset
 | 
2131  | 
in SOME (Vartab.fold (unify_vars o #2) vars env) end  | 
| 
 
0fa3b456a267
unify types of schematic variables in non-lifted case (i.e. "compose variants") -- allow schematic polymorphism, without revisiting HO-unification;
 
wenzelm 
parents: 
52131 
diff
changeset
 | 
2132  | 
handle Pattern.Unif => NONE;  | 
| 
 
0fa3b456a267
unify types of schematic variables in non-lifted case (i.e. "compose variants") -- allow schematic polymorphism, without revisiting HO-unification;
 
wenzelm 
parents: 
52131 
diff
changeset
 | 
2133  | 
|
| 0 | 2134  | 
(*Composition of object rule r=(A1...Am/B) with proof state s=(B1...Bn/C)  | 
| 250 | 2135  | 
Unifies B with Bi, replacing subgoal i (1 <= i <= n)  | 
| 0 | 2136  | 
If match then forbid instantiations in proof state  | 
2137  | 
If lifted then shorten the dpair using strip_assums2.  | 
|
2138  | 
If eres_flg then simultaneously proves A1 by assumption.  | 
|
| 250 | 2139  | 
nsubgoal is the number of new subgoals (written m above).  | 
| 0 | 2140  | 
Curried so that resolution calls dest_state only once.  | 
2141  | 
*)  | 
|
| 4270 | 2142  | 
local exception COMPOSE  | 
| 0 | 2143  | 
in  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
58946 
diff
changeset
 | 
2144  | 
fun bicompose_aux opt_ctxt {flatten, match, incremented} (state, (stpairs, Bs, Bi, C), lifted)
 | 
| 0 | 2145  | 
(eres_flg, orule, nsubgoal) =  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
2146  | 
 let val Thm (sder, {maxidx=smax, constraints = constraints2, shyps = shyps2, hyps = hyps2, ...}) = state
 | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
2147  | 
     and Thm (rder, {maxidx=rmax, constraints = constraints1, shyps = shyps1, hyps = hyps1,
 | 
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
2148  | 
tpairs=rtpairs, prop=rprop,...}) = orule  | 
| 1529 | 2149  | 
(*How many hyps to skip over during normalization*)  | 
| 21576 | 2150  | 
and nlift = Logic.count_prems (strip_all_body Bi) + (if eres_flg then ~1 else 0)  | 
| 
64981
 
ea6199b23dfa
proper background certificate from make_context, which can be a super-theory of the direct join (amending d07464875dd4);
 
wenzelm 
parents: 
64574 
diff
changeset
 | 
2151  | 
val (context, cert) =  | 
| 
 
ea6199b23dfa
proper background certificate from make_context, which can be a super-theory of the direct join (amending d07464875dd4);
 
wenzelm 
parents: 
64574 
diff
changeset
 | 
2152  | 
make_context_certificate [state, orule] opt_ctxt (join_certificate2 (state, orule));  | 
| 67721 | 2153  | 
(*Add new theorem with prop = "\<lbrakk>Bs; As\<rbrakk> \<Longrightarrow> C" to thq*)  | 
| 32032 | 2154  | 
fun addth A (As, oldAs, rder', n) ((env, tpairs), thq) =  | 
| 250 | 2155  | 
let val normt = Envir.norm_term env;  | 
2156  | 
(*perform minimal copying here by examining env*)  | 
|
| 
13658
 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 
berghofe 
parents: 
13642 
diff
changeset
 | 
2157  | 
val (ntpairs, normp) =  | 
| 
 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 
berghofe 
parents: 
13642 
diff
changeset
 | 
2158  | 
if Envir.is_empty env then (tpairs, (Bs @ As, C))  | 
| 250 | 2159  | 
else  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58950 
diff
changeset
 | 
2160  | 
let val ntps = map (apply2 normt) tpairs  | 
| 19861 | 2161  | 
in if Envir.above env smax then  | 
| 1238 | 2162  | 
(*no assignments in state; normalize the rule only*)  | 
2163  | 
if lifted  | 
|
| 
13658
 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 
berghofe 
parents: 
13642 
diff
changeset
 | 
2164  | 
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
 | 
2165  | 
else (ntps, (Bs @ map normt As, C))  | 
| 1529 | 2166  | 
else if match then raise COMPOSE  | 
| 250 | 2167  | 
else (*normalize the new rule fully*)  | 
| 
13658
 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 
berghofe 
parents: 
13642 
diff
changeset
 | 
2168  | 
(ntps, (map normt (Bs @ As), normt C))  | 
| 250 | 2169  | 
end  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
2170  | 
val constraints' =  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
2171  | 
union_constraints constraints1 constraints2  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
2172  | 
|> insert_constraints_env (Context.certificate_theory cert) env;  | 
| 
70822
 
22e2f5acc0b4
more accurate treatment of propositions within proof terms, but these are ultimately ignored for performance reasons;
 
wenzelm 
parents: 
70814 
diff
changeset
 | 
2173  | 
fun bicompose_proof prf1 prf2 =  | 
| 
 
22e2f5acc0b4
more accurate treatment of propositions within proof terms, but these are ultimately ignored for performance reasons;
 
wenzelm 
parents: 
70814 
diff
changeset
 | 
2174  | 
Proofterm.bicompose_proof flatten (map normt Bs) (map normt As) A oldAs n (nlift+1)  | 
| 
 
22e2f5acc0b4
more accurate treatment of propositions within proof terms, but these are ultimately ignored for performance reasons;
 
wenzelm 
parents: 
70814 
diff
changeset
 | 
2175  | 
prf1 prf2  | 
| 
16601
 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 
wenzelm 
parents: 
16425 
diff
changeset
 | 
2176  | 
val th =  | 
| 
52487
 
48bc24467008
backout dedd7952a62c: static "proofs" value within theory prevents later inferencing with different configuration;
 
wenzelm 
parents: 
52470 
diff
changeset
 | 
2177  | 
Thm (deriv_rule2  | 
| 70825 | 2178  | 
(if Envir.is_empty env then bicompose_proof  | 
2179  | 
else if Envir.above env smax then bicompose_proof o Proofterm.norm_proof' env  | 
|
2180  | 
else Proofterm.norm_proof' env oo bicompose_proof) rder' sder,  | 
|
| 
28321
 
9f4499bf9384
type thm: fully internal derivation, no longer exported;
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
2181  | 
                {tags = [],
 | 
| 32032 | 2182  | 
maxidx = Envir.maxidx_of env,  | 
| 
70459
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
2183  | 
constraints = constraints',  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
2184  | 
shyps = Envir.insert_sorts env (Sorts.union shyps1 shyps2),  | 
| 
 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 
wenzelm 
parents: 
70458 
diff
changeset
 | 
2185  | 
hyps = union_hyps hyps1 hyps2,  | 
| 
13658
 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 
berghofe 
parents: 
13642 
diff
changeset
 | 
2186  | 
tpairs = ntpairs,  | 
| 
24143
 
90a9a6fe0d01
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23781 
diff
changeset
 | 
2187  | 
prop = Logic.list_implies normp,  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
2188  | 
cert = cert})  | 
| 19475 | 2189  | 
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
 | 
2190  | 
val (rAs,B) = Logic.strip_prems(nsubgoal, [], rprop)  | 
| 0 | 2191  | 
       handle TERM _ => raise THM("bicompose: rule", 0, [orule,state]);
 | 
2192  | 
(*Modify assumptions, deleting n-th if n>0 for e-resolution*)  | 
|
2193  | 
fun newAs(As0, n, dpairs, tpairs) =  | 
|
| 11518 | 2194  | 
let val (As1, rder') =  | 
| 25939 | 2195  | 
if not lifted then (As0, rder)  | 
| 
44108
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
2196  | 
else  | 
| 
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
2197  | 
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
 | 
2198  | 
in (map (rename strip_apply) As0,  | 
| 
52487
 
48bc24467008
backout dedd7952a62c: static "proofs" value within theory prevents later inferencing with different configuration;
 
wenzelm 
parents: 
52470 
diff
changeset
 | 
2199  | 
deriv_rule1 (Proofterm.map_proof_terms (rename K) I) rder)  | 
| 
44108
 
476a239d3e0e
rename_bvs now avoids introducing name clashes between schematic variables
 
berghofe 
parents: 
43795 
diff
changeset
 | 
2200  | 
end;  | 
| 18486 | 2201  | 
in (map (if flatten then (Logic.flatten_params n) else I) As1, As1, rder', n)  | 
| 250 | 2202  | 
handle TERM _ =>  | 
2203  | 
          raise THM("bicompose: 1st premise", 0, [orule])
 | 
|
| 0 | 2204  | 
end;  | 
2205  | 
val BBi = if lifted then strip_assums2(B,Bi) else (B,Bi);  | 
|
2206  | 
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
 | 
2207  | 
|
| 
 
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
 | 
2208  | 
(*elim-resolution: try each assumption in turn*)  | 
| 
52222
 
0fa3b456a267
unify types of schematic variables in non-lifted case (i.e. "compose variants") -- allow schematic polymorphism, without revisiting HO-unification;
 
wenzelm 
parents: 
52131 
diff
changeset
 | 
2209  | 
     fun eres _ [] = raise THM ("bicompose: no premises", 0, [orule, state])
 | 
| 
 
0fa3b456a267
unify types of schematic variables in non-lifted case (i.e. "compose variants") -- allow schematic polymorphism, without revisiting HO-unification;
 
wenzelm 
parents: 
52131 
diff
changeset
 | 
2210  | 
| eres env (A1 :: As) =  | 
| 
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
 | 
2211  | 
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
 | 
2212  | 
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
 | 
2213  | 
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
 | 
2214  | 
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
 | 
2215  | 
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
 | 
2216  | 
| tryasms (asm :: rest) n =  | 
| 
 
7be15917f3fa
eq_assumption: slightly more efficient by checking (open) result of Logic.assum_problems directly;
 
wenzelm 
parents: 
30554 
diff
changeset
 | 
2217  | 
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
 | 
2218  | 
let val asm' = close asm in  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
58946 
diff
changeset
 | 
2219  | 
(case Seq.pull (Unify.unifiers (context, 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
 | 
2220  | 
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
 | 
2221  | 
| 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
 | 
2222  | 
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
 | 
2223  | 
(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
 | 
2224  | 
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
 | 
2225  | 
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
 | 
2226  | 
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
 | 
2227  | 
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
 | 
2228  | 
|
| 0 | 2229  | 
(*ordinary resolution*)  | 
| 
52222
 
0fa3b456a267
unify types of schematic variables in non-lifted case (i.e. "compose variants") -- allow schematic polymorphism, without revisiting HO-unification;
 
wenzelm 
parents: 
52131 
diff
changeset
 | 
2230  | 
fun res env =  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
58946 
diff
changeset
 | 
2231  | 
(case Seq.pull (Unify.unifiers (context, env, 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
 | 
2232  | 
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
 | 
2233  | 
| 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
 | 
2234  | 
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
 | 
2235  | 
(Seq.make (fn () => cell), Seq.empty));  | 
| 
52222
 
0fa3b456a267
unify types of schematic variables in non-lifted case (i.e. "compose variants") -- allow schematic polymorphism, without revisiting HO-unification;
 
wenzelm 
parents: 
52131 
diff
changeset
 | 
2236  | 
|
| 
 
0fa3b456a267
unify types of schematic variables in non-lifted case (i.e. "compose variants") -- allow schematic polymorphism, without revisiting HO-unification;
 
wenzelm 
parents: 
52131 
diff
changeset
 | 
2237  | 
val env0 = Envir.empty (Int.max (rmax, smax));  | 
| 
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
 | 
2238  | 
in  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
58946 
diff
changeset
 | 
2239  | 
(case if incremented then SOME env0 else unify_var_types context (state, orule) env0 of  | 
| 
52222
 
0fa3b456a267
unify types of schematic variables in non-lifted case (i.e. "compose variants") -- allow schematic polymorphism, without revisiting HO-unification;
 
wenzelm 
parents: 
52131 
diff
changeset
 | 
2240  | 
NONE => Seq.empty  | 
| 
 
0fa3b456a267
unify types of schematic variables in non-lifted case (i.e. "compose variants") -- allow schematic polymorphism, without revisiting HO-unification;
 
wenzelm 
parents: 
52131 
diff
changeset
 | 
2241  | 
| SOME env => if eres_flg then eres env (rev rAs) else res env)  | 
| 0 | 2242  | 
end;  | 
| 7528 | 2243  | 
end;  | 
| 0 | 2244  | 
|
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
58946 
diff
changeset
 | 
2245  | 
fun bicompose opt_ctxt flags arg i state =  | 
| 
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
58946 
diff
changeset
 | 
2246  | 
bicompose_aux opt_ctxt flags (state, dest_state (state,i), false) arg;  | 
| 0 | 2247  | 
|
2248  | 
(*Quick test whether rule is resolvable with the subgoal with hyps Hs  | 
|
2249  | 
and conclusion B. If eres_flg then checks 1st premise of rule also*)  | 
|
2250  | 
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
 | 
2251  | 
let fun could_reshyp (A1::_) = exists (fn H => Term.could_unify (A1, H)) Hs  | 
| 250 | 2252  | 
| 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
 | 
2253  | 
in Term.could_unify(concl_of rule, B) andalso  | 
| 250 | 2254  | 
(not eres_flg orelse could_reshyp (prems_of rule))  | 
| 0 | 2255  | 
end;  | 
2256  | 
||
2257  | 
(*Bi-resolution of a state with a list of (flag,rule) pairs.  | 
|
2258  | 
Puts the rule above: rule/state. Renames vars in the rules. *)  | 
|
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
58946 
diff
changeset
 | 
2259  | 
fun biresolution opt_ctxt match brules i state =  | 
| 18035 | 2260  | 
let val (stpairs, Bs, Bi, C) = dest_state(state,i);  | 
| 18145 | 2261  | 
val lift = lift_rule (cprem_of state i);  | 
| 250 | 2262  | 
val B = Logic.strip_assums_concl Bi;  | 
2263  | 
val Hs = Logic.strip_assums_hyp Bi;  | 
|
| 
52223
 
5bb6ae8acb87
tuned signature -- more explicit flags for low-level Thm.bicompose;
 
wenzelm 
parents: 
52222 
diff
changeset
 | 
2264  | 
val compose =  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
58946 
diff
changeset
 | 
2265  | 
          bicompose_aux opt_ctxt {flatten = true, match = match, incremented = true}
 | 
| 
52223
 
5bb6ae8acb87
tuned signature -- more explicit flags for low-level Thm.bicompose;
 
wenzelm 
parents: 
52222 
diff
changeset
 | 
2266  | 
(state, (stpairs, Bs, Bi, C), true);  | 
| 4270 | 2267  | 
fun res [] = Seq.empty  | 
| 250 | 2268  | 
| res ((eres_flg, rule)::brules) =  | 
| 
61044
 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 
wenzelm 
parents: 
61042 
diff
changeset
 | 
2269  | 
if Config.get_generic (make_context [state] opt_ctxt (cert_of state))  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
58946 
diff
changeset
 | 
2270  | 
Pattern.unify_trace_failure orelse could_bires (Hs, B, eres_flg, rule)  | 
| 4270 | 2271  | 
then Seq.make (*delay processing remainder till needed*)  | 
| 22573 | 2272  | 
(fn()=> SOME(compose (eres_flg, lift rule, nprems_of rule),  | 
| 250 | 2273  | 
res brules))  | 
2274  | 
else res brules  | 
|
| 4270 | 2275  | 
in Seq.flat (res brules) end;  | 
| 0 | 2276  | 
|
| 
70456
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2277  | 
(*Resolution: exactly one resolvent must be produced*)  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2278  | 
fun tha RSN (i, thb) =  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2279  | 
(case Seq.chop 2 (biresolution NONE false [(false, tha)] i thb) of  | 
| 
70472
 
cf66d2db97fe
more robust and convenient treatment of implicit context;
 
wenzelm 
parents: 
70469 
diff
changeset
 | 
2280  | 
([th], _) => solve_constraints th  | 
| 
70456
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2281  | 
  | ([], _) => raise THM ("RSN: no unifiers", i, [tha, thb])
 | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2282  | 
  | _ => raise THM ("RSN: multiple unifiers", i, [tha, thb]));
 | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2283  | 
|
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2284  | 
(*Resolution: P \<Longrightarrow> Q, Q \<Longrightarrow> R gives P \<Longrightarrow> R*)  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2285  | 
fun tha RS thb = tha RSN (1,thb);  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2286  | 
|
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2287  | 
|
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2288  | 
|
| 70458 | 2289  | 
(**** Type classes ****)  | 
| 
70456
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2290  | 
|
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2291  | 
fun standard_tvars thm =  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2292  | 
let  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2293  | 
val thy = theory_of_thm thm;  | 
| 74235 | 2294  | 
val tvars = build_rev (Term.add_tvars (prop_of thm));  | 
| 
70456
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2295  | 
val names = Name.invent Name.context Name.aT (length tvars);  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2296  | 
val tinst =  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2297  | 
map2 (fn (ai, S) => fn b => ((ai, S), global_ctyp_of thy (TVar ((b, 0), S)))) tvars names;  | 
| 74282 | 2298  | 
in instantiate (TVars.make tinst, Vars.empty) thm end  | 
| 
70456
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2299  | 
|
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2300  | 
|
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2301  | 
(* class relations *)  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2302  | 
|
| 70457 | 2303  | 
val is_classrel = Symreltab.defined o get_classrels;  | 
2304  | 
||
| 
70456
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2305  | 
fun complete_classrels thy =  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2306  | 
let  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2307  | 
fun complete (c, (_, (all_preds, all_succs))) (finished1, thy1) =  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2308  | 
let  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2309  | 
fun compl c1 c2 (finished2, thy2) =  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2310  | 
if is_classrel thy2 (c1, c2) then (finished2, thy2)  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2311  | 
else  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2312  | 
(false,  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2313  | 
thy2  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2314  | 
|> (map_classrels o Symreltab.update) ((c1, c2),  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2315  | 
(the_classrel thy2 (c1, c) RS the_classrel thy2 (c, c2))  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2316  | 
|> standard_tvars  | 
| 
71018
 
d32ed8927a42
more robust expose_proofs corresponding to register_proofs/consolidate_theory;
 
wenzelm 
parents: 
71017 
diff
changeset
 | 
2317  | 
|> close_derivation \<^here>  | 
| 
 
d32ed8927a42
more robust expose_proofs corresponding to register_proofs/consolidate_theory;
 
wenzelm 
parents: 
71017 
diff
changeset
 | 
2318  | 
|> tap (expose_proof thy2)  | 
| 
70456
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2319  | 
|> trim_context));  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2320  | 
|
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2321  | 
val proven = is_classrel thy1;  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2322  | 
val preds = Graph.Keys.fold (fn c1 => proven (c1, c) ? cons c1) all_preds [];  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2323  | 
val succs = Graph.Keys.fold (fn c2 => proven (c, c2) ? cons c2) all_succs [];  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2324  | 
in  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2325  | 
fold_product compl preds succs (finished1, thy1)  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2326  | 
end;  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2327  | 
in  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2328  | 
(case Graph.fold complete (Sorts.classes_of (Sign.classes_of thy)) (true, thy) of  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2329  | 
(true, _) => NONE  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2330  | 
| (_, thy') => SOME thy')  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2331  | 
end;  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2332  | 
|
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2333  | 
|
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2334  | 
(* type arities *)  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2335  | 
|
| 
70475
 
98b6da301e13
backed out changeset 1b8858f4c393: odd problems e.g. in CAVA_LTL_Modelchecker;
 
wenzelm 
parents: 
70472 
diff
changeset
 | 
2336  | 
fun thynames_of_arity thy (a, c) =  | 
| 74234 | 2337  | 
build (get_arities thy |> Aritytab.fold  | 
2338  | 
(fn ((a', _, c'), (_, name, ser)) => (a = a' andalso c = c') ? cons (name, ser)))  | 
|
| 
70475
 
98b6da301e13
backed out changeset 1b8858f4c393: odd problems e.g. in CAVA_LTL_Modelchecker;
 
wenzelm 
parents: 
70472 
diff
changeset
 | 
2339  | 
|> sort (int_ord o apply2 #2) |> map #1;  | 
| 
 
98b6da301e13
backed out changeset 1b8858f4c393: odd problems e.g. in CAVA_LTL_Modelchecker;
 
wenzelm 
parents: 
70472 
diff
changeset
 | 
2340  | 
|
| 
 
98b6da301e13
backed out changeset 1b8858f4c393: odd problems e.g. in CAVA_LTL_Modelchecker;
 
wenzelm 
parents: 
70472 
diff
changeset
 | 
2341  | 
fun insert_arity_completions thy ((t, Ss, c), (th, thy_name, ser)) (finished, arities) =  | 
| 
70456
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2342  | 
let  | 
| 70461 | 2343  | 
val completions =  | 
2344  | 
Sign.super_classes thy c |> map_filter (fn c1 =>  | 
|
| 70465 | 2345  | 
if Aritytab.defined arities (t, Ss, c1) then NONE  | 
| 70461 | 2346  | 
else  | 
2347  | 
let  | 
|
2348  | 
val th1 =  | 
|
2349  | 
(th RS the_classrel thy (c, c1))  | 
|
2350  | 
|> standard_tvars  | 
|
| 
71018
 
d32ed8927a42
more robust expose_proofs corresponding to register_proofs/consolidate_theory;
 
wenzelm 
parents: 
71017 
diff
changeset
 | 
2351  | 
|> close_derivation \<^here>  | 
| 
 
d32ed8927a42
more robust expose_proofs corresponding to register_proofs/consolidate_theory;
 
wenzelm 
parents: 
71017 
diff
changeset
 | 
2352  | 
|> tap (expose_proof thy)  | 
| 70461 | 2353  | 
|> trim_context;  | 
| 
70475
 
98b6da301e13
backed out changeset 1b8858f4c393: odd problems e.g. in CAVA_LTL_Modelchecker;
 
wenzelm 
parents: 
70472 
diff
changeset
 | 
2354  | 
in SOME ((t, Ss, c1), (th1, thy_name, ser)) end);  | 
| 
70456
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2355  | 
val finished' = finished andalso null completions;  | 
| 70465 | 2356  | 
val arities' = fold Aritytab.update completions arities;  | 
| 
70456
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2357  | 
in (finished', arities') end;  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2358  | 
|
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2359  | 
fun complete_arities thy =  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2360  | 
let  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2361  | 
val arities = get_arities thy;  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2362  | 
val (finished, arities') =  | 
| 70465 | 2363  | 
Aritytab.fold (insert_arity_completions thy) arities (true, get_arities thy);  | 
| 
70456
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2364  | 
in  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2365  | 
if finished then NONE  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2366  | 
else SOME (map_arities (K arities') thy)  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2367  | 
end;  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2368  | 
|
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2369  | 
val _ =  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2370  | 
Theory.setup  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2371  | 
(Theory.at_begin complete_classrels #>  | 
| 70458 | 2372  | 
Theory.at_begin complete_arities);  | 
| 
70456
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2373  | 
|
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2374  | 
|
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2375  | 
(* primitive rules *)  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2376  | 
|
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2377  | 
fun add_classrel raw_th thy =  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2378  | 
let  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2379  | 
val th = strip_shyps (transfer thy raw_th);  | 
| 
71017
 
c85efa2be619
expose derivations more thoroughly, notably for locale/class reasoning;
 
wenzelm 
parents: 
71012 
diff
changeset
 | 
2380  | 
val th' = th |> unconstrainT |> tap (expose_proof thy) |> trim_context;  | 
| 
70456
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2381  | 
val prop = plain_prop_of th;  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2382  | 
val (c1, c2) = Logic.dest_classrel prop;  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2383  | 
in  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2384  | 
thy  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2385  | 
|> Sign.primitive_classrel (c1, c2)  | 
| 
71017
 
c85efa2be619
expose derivations more thoroughly, notably for locale/class reasoning;
 
wenzelm 
parents: 
71012 
diff
changeset
 | 
2386  | 
|> map_classrels (Symreltab.update ((c1, c2), th'))  | 
| 
70456
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2387  | 
|> perhaps complete_classrels  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2388  | 
|> perhaps complete_arities  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2389  | 
end;  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2390  | 
|
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2391  | 
fun add_arity raw_th thy =  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2392  | 
let  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2393  | 
val th = strip_shyps (transfer thy raw_th);  | 
| 
71017
 
c85efa2be619
expose derivations more thoroughly, notably for locale/class reasoning;
 
wenzelm 
parents: 
71012 
diff
changeset
 | 
2394  | 
val th' = th |> unconstrainT |> tap (expose_proof thy) |> trim_context;  | 
| 
70456
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2395  | 
val prop = plain_prop_of th;  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2396  | 
val (t, Ss, c) = Logic.dest_arity prop;  | 
| 
71017
 
c85efa2be619
expose derivations more thoroughly, notably for locale/class reasoning;
 
wenzelm 
parents: 
71012 
diff
changeset
 | 
2397  | 
val ar = ((t, Ss, c), (th', Context.theory_name thy, serial ()));  | 
| 
70456
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2398  | 
in  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2399  | 
thy  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2400  | 
|> Sign.primitive_arity (t, Ss, [c])  | 
| 70465 | 2401  | 
|> map_arities (Aritytab.update ar #> curry (insert_arity_completions thy ar) true #> #2)  | 
| 
70456
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2402  | 
end;  | 
| 
 
c742527a25fe
clarified modules: inference kernel maintains sort algebra within the logic;
 
wenzelm 
parents: 
70454 
diff
changeset
 | 
2403  | 
|
| 0 | 2404  | 
end;  | 
| 1503 | 2405  | 
|
| 32104 | 2406  | 
structure Basic_Thm: BASIC_THM = Thm;  | 
2407  | 
open Basic_Thm;  |