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