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