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