| author | huffman | 
| Wed, 09 May 2007 00:33:12 +0200 | |
| changeset 22881 | c23ded11158f | 
| parent 22685 | fc4ef3807fb9 | 
| child 22909 | 7de3b0ac4189 | 
| permissions | -rw-r--r-- | 
| 250 | 1 | (* Title: Pure/thm.ML | 
| 0 | 2 | ID: $Id$ | 
| 250 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
225diff
changeset | 4 | Copyright 1994 University of Cambridge | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
225diff
changeset | 5 | |
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 6 | The very core of Isabelle's Meta Logic: certified types and terms, | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 7 | meta theorems, meta rules (including lifting and resolution). | 
| 0 | 8 | *) | 
| 9 | ||
| 6089 | 10 | signature BASIC_THM = | 
| 1503 | 11 | sig | 
| 1160 | 12 | (*certified types*) | 
| 387 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 13 | type ctyp | 
| 16656 | 14 | val rep_ctyp: ctyp -> | 
| 15 |    {thy: theory,
 | |
| 16 | T: typ, | |
| 20512 | 17 | maxidx: int, | 
| 16656 | 18 | sorts: sort list} | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 19 | val theory_of_ctyp: ctyp -> theory | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 20 | val typ_of: ctyp -> typ | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 21 | val ctyp_of: theory -> typ -> ctyp | 
| 1160 | 22 | |
| 23 | (*certified terms*) | |
| 24 | type cterm | |
| 22584 | 25 | exception CTERM of string * cterm list | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 26 | val rep_cterm: cterm -> | 
| 16656 | 27 |    {thy: theory,
 | 
| 28 | t: term, | |
| 29 | T: typ, | |
| 30 | maxidx: int, | |
| 31 | sorts: sort list} | |
| 22596 | 32 |   val crep_cterm: cterm -> {thy: theory, t: term, T: ctyp, maxidx: int, sorts: sort list}
 | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 33 | val theory_of_cterm: cterm -> theory | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 34 | val term_of: cterm -> term | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 35 | val cterm_of: theory -> term -> cterm | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 36 | val ctyp_of_term: cterm -> ctyp | 
| 1160 | 37 | |
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 38 | type tag (* = string * string list *) | 
| 1529 | 39 | |
| 1160 | 40 | (*meta theorems*) | 
| 41 | type thm | |
| 22365 | 42 | type attribute (* = Context.generic * thm -> Context.generic * thm *) | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 43 | val rep_thm: thm -> | 
| 16656 | 44 |    {thy: theory,
 | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 45 | der: bool * Proofterm.proof, | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 46 | tags: tag list, | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 47 | maxidx: int, | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 48 | shyps: sort list, | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 49 | hyps: term list, | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 50 | tpairs: (term * term) list, | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 51 | prop: term} | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 52 | val crep_thm: thm -> | 
| 16656 | 53 |    {thy: theory,
 | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 54 | der: bool * Proofterm.proof, | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 55 | tags: tag list, | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 56 | maxidx: int, | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 57 | shyps: sort list, | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 58 | hyps: cterm list, | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 59 | tpairs: (cterm * cterm) list, | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 60 | prop: cterm} | 
| 6089 | 61 | exception THM of string * int * thm list | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 62 | val theory_of_thm: thm -> theory | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 63 | val prop_of: thm -> term | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 64 | val proof_of: thm -> Proofterm.proof | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 65 | val tpairs_of: thm -> (term * term) list | 
| 16656 | 66 | val concl_of: thm -> term | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 67 | val prems_of: thm -> term list | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 68 | val nprems_of: thm -> int | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 69 | val cprop_of: thm -> cterm | 
| 18145 | 70 | val cprem_of: thm -> int -> cterm | 
| 16656 | 71 | val transfer: theory -> thm -> thm | 
| 16945 | 72 | val weaken: cterm -> thm -> thm | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 73 | val extra_shyps: thm -> sort list | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 74 | val strip_shyps: thm -> thm | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 75 | val get_axiom_i: theory -> string -> thm | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 76 | val get_axiom: theory -> xstring -> thm | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 77 | val def_name: string -> string | 
| 20884 | 78 | val def_name_optional: string -> string -> string | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 79 | val get_def: theory -> xstring -> thm | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 80 | val axioms_of: theory -> (string * thm) list | 
| 1160 | 81 | |
| 82 | (*meta rules*) | |
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 83 | val assume: cterm -> thm | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 84 | val implies_intr: cterm -> thm -> thm | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 85 | val implies_elim: thm -> thm -> thm | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 86 | val forall_intr: cterm -> thm -> thm | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 87 | val forall_elim: cterm -> thm -> thm | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 88 | val reflexive: cterm -> thm | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 89 | val symmetric: thm -> thm | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 90 | val transitive: thm -> thm -> thm | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 91 | val beta_conversion: bool -> cterm -> thm | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 92 | val eta_conversion: cterm -> thm | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 93 | val abstract_rule: string -> cterm -> thm -> thm | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 94 | val combination: thm -> thm -> thm | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 95 | val equal_intr: thm -> thm -> thm | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 96 | val equal_elim: thm -> thm -> thm | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 97 | val flexflex_rule: thm -> thm Seq.seq | 
| 19910 | 98 | val generalize: string list * string list -> int -> thm -> thm | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 99 | val instantiate: (ctyp * ctyp) list * (cterm * cterm) list -> thm -> thm | 
| 22584 | 100 | val instantiate_cterm: (ctyp * ctyp) list * (cterm * cterm) list -> cterm -> cterm | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 101 | val trivial: cterm -> thm | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 102 | val class_triv: theory -> class -> thm | 
| 19505 | 103 | val unconstrainT: ctyp -> thm -> thm | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 104 | val dest_state: thm * int -> (term * term) list * term list * term * term | 
| 18035 | 105 | val lift_rule: cterm -> thm -> thm | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 106 | val incr_indexes: int -> thm -> thm | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 107 | val assumption: int -> thm -> thm Seq.seq | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 108 | val eq_assumption: int -> thm -> thm | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 109 | val rotate_rule: int -> int -> thm -> thm | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 110 | val permute_prems: int -> int -> thm -> thm | 
| 1160 | 111 | val rename_params_rule: string list * int -> thm -> thm | 
| 18501 
915105af2e80
turned bicompose_no_flatten into compose_no_flatten, without elimination;
 wenzelm parents: 
18486diff
changeset | 112 | val compose_no_flatten: bool -> thm * int -> int -> thm -> thm Seq.seq | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 113 | val bicompose: bool -> bool * thm * int -> int -> thm -> thm Seq.seq | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 114 | val biresolution: bool -> (bool * thm) list -> int -> thm -> thm Seq.seq | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 115 | val invoke_oracle: theory -> xstring -> theory * Object.T -> thm | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 116 | val invoke_oracle_i: theory -> string -> theory * Object.T -> thm | 
| 250 | 117 | end; | 
| 0 | 118 | |
| 6089 | 119 | signature THM = | 
| 120 | sig | |
| 121 | include BASIC_THM | |
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 122 | val dest_ctyp: ctyp -> ctyp list | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 123 | val dest_comb: cterm -> cterm * cterm | 
| 20580 
6fb75df09253
added dest_arg, i.e. a tuned version of #2 o dest_comb;
 wenzelm parents: 
20548diff
changeset | 124 | val dest_arg: cterm -> cterm | 
| 20673 | 125 | val dest_binop: cterm -> cterm * cterm | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 126 | val dest_abs: string option -> cterm -> cterm * cterm | 
| 20261 | 127 | val adjust_maxidx_cterm: int -> cterm -> cterm | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 128 | val capply: cterm -> cterm -> cterm | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 129 | val cabs: cterm -> cterm -> cterm | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 130 | val major_prem_of: thm -> term | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 131 | val no_prems: thm -> bool | 
| 16945 | 132 | val terms_of_tpairs: (term * term) list -> term list | 
| 19881 | 133 | val maxidx_of: thm -> int | 
| 19910 | 134 | val maxidx_thm: thm -> int -> int | 
| 19881 | 135 | val hyps_of: thm -> term list | 
| 16945 | 136 | val full_prop_of: thm -> term | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 137 | val get_name: thm -> string | 
| 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 138 | val put_name: string -> thm -> thm | 
| 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 139 | val get_tags: thm -> tag list | 
| 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 140 | val map_tags: (tag list -> tag list) -> thm -> thm | 
| 16945 | 141 | val compress: thm -> thm | 
| 20261 | 142 | val adjust_maxidx_thm: int -> thm -> thm | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 143 | val rename_boundvars: term -> term -> thm -> thm | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 144 | val cterm_match: cterm * cterm -> (ctyp * ctyp) list * (cterm * cterm) list | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 145 | val cterm_first_order_match: cterm * cterm -> (ctyp * ctyp) list * (cterm * cterm) list | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 146 | val cterm_incr_indexes: int -> cterm -> cterm | 
| 20002 | 147 | val varifyT: thm -> thm | 
| 148 | val varifyT': (string * sort) list -> thm -> ((string * sort) * indexname) list * thm | |
| 19881 | 149 | val freezeT: thm -> thm | 
| 6089 | 150 | end; | 
| 151 | ||
| 3550 | 152 | structure Thm: THM = | 
| 0 | 153 | struct | 
| 250 | 154 | |
| 22237 
bb9b1c8a8a95
old-fashioned abstype ctyp/cterm/thm prevents accidental equality tests;
 wenzelm parents: 
21975diff
changeset | 155 | structure Pt = Proofterm; | 
| 
bb9b1c8a8a95
old-fashioned abstype ctyp/cterm/thm prevents accidental equality tests;
 wenzelm parents: 
21975diff
changeset | 156 | |
| 16656 | 157 | |
| 387 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 158 | (*** Certified terms and types ***) | 
| 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 159 | |
| 16656 | 160 | (** collect occurrences of sorts -- unless all sorts non-empty **) | 
| 161 | ||
| 16679 | 162 | fun may_insert_typ_sorts thy T = if Sign.all_sorts_nonempty thy then I else Sorts.insert_typ T; | 
| 163 | fun may_insert_term_sorts thy t = if Sign.all_sorts_nonempty thy then I else Sorts.insert_term t; | |
| 16656 | 164 | |
| 165 | (*NB: type unification may invent new sorts*) | |
| 166 | fun may_insert_env_sorts thy (env as Envir.Envir {iTs, ...}) =
 | |
| 167 | if Sign.all_sorts_nonempty thy then I | |
| 168 | else Vartab.fold (fn (_, (_, T)) => Sorts.insert_typ T) iTs; | |
| 169 | ||
| 170 | ||
| 171 | ||
| 250 | 172 | (** certified types **) | 
| 173 | ||
| 22237 
bb9b1c8a8a95
old-fashioned abstype ctyp/cterm/thm prevents accidental equality tests;
 wenzelm parents: 
21975diff
changeset | 174 | abstype ctyp = Ctyp of | 
| 20512 | 175 |  {thy_ref: theory_ref,
 | 
| 176 | T: typ, | |
| 177 | maxidx: int, | |
| 22237 
bb9b1c8a8a95
old-fashioned abstype ctyp/cterm/thm prevents accidental equality tests;
 wenzelm parents: 
21975diff
changeset | 178 | sorts: sort list} | 
| 
bb9b1c8a8a95
old-fashioned abstype ctyp/cterm/thm prevents accidental equality tests;
 wenzelm parents: 
21975diff
changeset | 179 | with | 
| 250 | 180 | |
| 20512 | 181 | fun rep_ctyp (Ctyp {thy_ref, T, maxidx, sorts}) =
 | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 182 | let val thy = Theory.deref thy_ref | 
| 22596 | 183 |   in {thy = thy, T = T, maxidx = maxidx, sorts = sorts} end;
 | 
| 250 | 184 | |
| 16656 | 185 | fun theory_of_ctyp (Ctyp {thy_ref, ...}) = Theory.deref thy_ref;
 | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 186 | |
| 250 | 187 | fun typ_of (Ctyp {T, ...}) = T;
 | 
| 188 | ||
| 16656 | 189 | fun ctyp_of thy raw_T = | 
| 20512 | 190 | let val T = Sign.certify_typ thy raw_T in | 
| 191 |     Ctyp {thy_ref = Theory.self_ref thy, T = T,
 | |
| 192 | maxidx = Term.maxidx_of_typ T, sorts = may_insert_typ_sorts thy T []} | |
| 193 | end; | |
| 250 | 194 | |
| 20512 | 195 | fun dest_ctyp (Ctyp {thy_ref, T = Type (s, Ts), maxidx, sorts}) =
 | 
| 196 |       map (fn T => Ctyp {thy_ref = thy_ref, T = T, maxidx = maxidx, sorts = sorts}) Ts
 | |
| 16679 | 197 |   | dest_ctyp cT = raise TYPE ("dest_ctyp", [typ_of cT], []);
 | 
| 15087 | 198 | |
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
225diff
changeset | 199 | |
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
225diff
changeset | 200 | |
| 250 | 201 | (** certified terms **) | 
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
225diff
changeset | 202 | |
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 203 | (*certified terms with checked typ, maxidx, and sorts*) | 
| 22237 
bb9b1c8a8a95
old-fashioned abstype ctyp/cterm/thm prevents accidental equality tests;
 wenzelm parents: 
21975diff
changeset | 204 | abstype cterm = Cterm of | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 205 |  {thy_ref: theory_ref,
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 206 | t: term, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 207 | T: typ, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 208 | maxidx: int, | 
| 22237 
bb9b1c8a8a95
old-fashioned abstype ctyp/cterm/thm prevents accidental equality tests;
 wenzelm parents: 
21975diff
changeset | 209 | sorts: sort list} | 
| 
bb9b1c8a8a95
old-fashioned abstype ctyp/cterm/thm prevents accidental equality tests;
 wenzelm parents: 
21975diff
changeset | 210 | with | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 211 | |
| 22584 | 212 | exception CTERM of string * cterm list; | 
| 16679 | 213 | |
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 214 | fun rep_cterm (Cterm {thy_ref, t, T, maxidx, sorts}) =
 | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 215 | let val thy = Theory.deref thy_ref | 
| 22596 | 216 |   in {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 | 217 | |
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 218 | fun crep_cterm (Cterm {thy_ref, t, T, maxidx, sorts}) =
 | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 219 | let val thy = Theory.deref thy_ref in | 
| 22596 | 220 |    {thy = thy, t = t,
 | 
| 20512 | 221 |       T = Ctyp {thy_ref = thy_ref, T = T, maxidx = maxidx, sorts = sorts},
 | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 222 | maxidx = maxidx, sorts = sorts} | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 223 | end; | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 224 | |
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 225 | fun theory_of_cterm (Cterm {thy_ref, ...}) = Theory.deref thy_ref;
 | 
| 250 | 226 | fun term_of (Cterm {t, ...}) = t;
 | 
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
225diff
changeset | 227 | |
| 20512 | 228 | fun ctyp_of_term (Cterm {thy_ref, T, maxidx, sorts, ...}) =
 | 
| 229 |   Ctyp {thy_ref = thy_ref, T = T, maxidx = maxidx, sorts = sorts};
 | |
| 2671 | 230 | |
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 231 | fun cterm_of thy tm = | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 232 | let | 
| 18969 | 233 | val (t, T, maxidx) = Sign.certify_term thy tm; | 
| 16656 | 234 | val sorts = may_insert_term_sorts thy t []; | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 235 |   in Cterm {thy_ref = Theory.self_ref 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 | 236 | |
| 20057 | 237 | fun merge_thys0 (Cterm {thy_ref = r1, t = t1, ...}) (Cterm {thy_ref = r2, t = t2, ...}) =
 | 
| 238 | Theory.merge_refs (r1, r2) handle TERM (msg, _) => raise TERM (msg, [t1, t2]); | |
| 16656 | 239 | |
| 20580 
6fb75df09253
added dest_arg, i.e. a tuned version of #2 o dest_comb;
 wenzelm parents: 
20548diff
changeset | 240 | |
| 22584 | 241 | fun dest_comb (ct as Cterm {t = t $ u, T, thy_ref, maxidx, sorts}) =
 | 
| 16679 | 242 | let val A = Term.argument_type_of t in | 
| 243 |         (Cterm {t = t, T = A --> T, thy_ref = thy_ref, maxidx = maxidx, sorts = sorts},
 | |
| 244 |          Cterm {t = u, T = A, thy_ref = thy_ref, maxidx = maxidx, sorts = sorts})
 | |
| 1493 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 clasohm parents: 
1460diff
changeset | 245 | end | 
| 22584 | 246 |   | 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 | 247 | |
| 22584 | 248 | fun dest_arg (ct as Cterm {t = t $ u, T, thy_ref, maxidx, sorts}) =
 | 
| 20580 
6fb75df09253
added dest_arg, i.e. a tuned version of #2 o dest_comb;
 wenzelm parents: 
20548diff
changeset | 249 | let val A = Term.argument_type_of t in | 
| 
6fb75df09253
added dest_arg, i.e. a tuned version of #2 o dest_comb;
 wenzelm parents: 
20548diff
changeset | 250 |          Cterm {t = u, T = A, thy_ref = thy_ref, maxidx = maxidx, sorts = sorts}
 | 
| 
6fb75df09253
added dest_arg, i.e. a tuned version of #2 o dest_comb;
 wenzelm parents: 
20548diff
changeset | 251 | end | 
| 22584 | 252 |   | 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 | 253 | |
| 22584 | 254 | fun dest_binop (ct as Cterm {t = tm, T = _, thy_ref, maxidx, sorts}) =
 | 
| 20673 | 255 |   let fun cterm t T = Cterm {t = t, T = T, thy_ref = thy_ref, maxidx = maxidx, sorts = sorts} in
 | 
| 256 | (case tm of | |
| 257 |       Const (_, Type ("fun", [A, Type ("fun", [B, _])])) $ a $ b => (cterm a A, cterm b B)
 | |
| 258 |     |  Free (_, Type ("fun", [A, Type ("fun", [B, _])])) $ a $ b => (cterm a A, cterm b B)
 | |
| 259 |     |   Var (_, Type ("fun", [A, Type ("fun", [B, _])])) $ a $ b => (cterm a A, cterm b B)
 | |
| 22584 | 260 |     | _ => raise CTERM ("dest_binop", [ct]))
 | 
| 20673 | 261 | end; | 
| 262 | ||
| 22584 | 263 | fun dest_abs a (ct as | 
| 264 |         Cterm {t = Abs (x, T, t), T = Type ("fun", [_, U]), thy_ref, maxidx, sorts}) =
 | |
| 18944 | 265 | let val (y', t') = Term.dest_abs (the_default x a, T, t) in | 
| 16679 | 266 |         (Cterm {t = Free (y', T), T = T, thy_ref = thy_ref, maxidx = maxidx, sorts = sorts},
 | 
| 267 |           Cterm {t = t', T = U, thy_ref = thy_ref, maxidx = maxidx, sorts = sorts})
 | |
| 1493 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 clasohm parents: 
1460diff
changeset | 268 | end | 
| 22584 | 269 |   | 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 | 270 | |
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 271 | fun capply | 
| 16656 | 272 |   (cf as Cterm {t = f, T = Type ("fun", [dty, rty]), maxidx = maxidx1, sorts = sorts1, ...})
 | 
| 273 |   (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 | 274 | if T = dty then | 
| 16656 | 275 |       Cterm {thy_ref = merge_thys0 cf cx,
 | 
| 276 | t = f $ x, | |
| 277 | T = rty, | |
| 278 | maxidx = Int.max (maxidx1, maxidx2), | |
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 279 | sorts = Sorts.union sorts1 sorts2} | 
| 22584 | 280 |       else raise CTERM ("capply: types don't agree", [cf, cx])
 | 
| 281 |   | capply cf cx = raise CTERM ("capply: first arg is not a function", [cf, cx]);
 | |
| 250 | 282 | |
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 283 | fun cabs | 
| 16656 | 284 |   (ct1 as Cterm {t = t1, T = T1, maxidx = maxidx1, sorts = sorts1, ...})
 | 
| 285 |   (ct2 as Cterm {t = t2, T = T2, maxidx = maxidx2, sorts = sorts2, ...}) =
 | |
| 21975 
1152dc45d591
Term.lambda: abstract over arbitrary closed terms;
 wenzelm parents: 
21798diff
changeset | 286 | let val t = Term.lambda t1 t2 in | 
| 16656 | 287 |       Cterm {thy_ref = merge_thys0 ct1 ct2,
 | 
| 288 | t = t, T = T1 --> T2, | |
| 289 | maxidx = Int.max (maxidx1, maxidx2), | |
| 290 | sorts = Sorts.union sorts1 sorts2} | |
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 291 | end; | 
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
225diff
changeset | 292 | |
| 20580 
6fb75df09253
added dest_arg, i.e. a tuned version of #2 o dest_comb;
 wenzelm parents: 
20548diff
changeset | 293 | |
| 
6fb75df09253
added dest_arg, i.e. a tuned version of #2 o dest_comb;
 wenzelm parents: 
20548diff
changeset | 294 | fun adjust_maxidx_cterm i (ct as Cterm {thy_ref, t, T, maxidx, sorts}) =
 | 
| 
6fb75df09253
added dest_arg, i.e. a tuned version of #2 o dest_comb;
 wenzelm parents: 
20548diff
changeset | 295 | if maxidx = i then ct | 
| 
6fb75df09253
added dest_arg, i.e. a tuned version of #2 o dest_comb;
 wenzelm parents: 
20548diff
changeset | 296 | else if maxidx < i then | 
| 
6fb75df09253
added dest_arg, i.e. a tuned version of #2 o dest_comb;
 wenzelm parents: 
20548diff
changeset | 297 |     Cterm {maxidx = i, thy_ref = thy_ref, t = t, T = T, sorts = sorts}
 | 
| 
6fb75df09253
added dest_arg, i.e. a tuned version of #2 o dest_comb;
 wenzelm parents: 
20548diff
changeset | 298 | else | 
| 
6fb75df09253
added dest_arg, i.e. a tuned version of #2 o dest_comb;
 wenzelm parents: 
20548diff
changeset | 299 |     Cterm {maxidx = Int.max (maxidx_of_term t, i), thy_ref = thy_ref, t = t, T = T, sorts = sorts};
 | 
| 
6fb75df09253
added dest_arg, i.e. a tuned version of #2 o dest_comb;
 wenzelm parents: 
20548diff
changeset | 300 | |
| 10416 
5b33e732e459
- Moved rewriting functions to meta_simplifier.ML
 berghofe parents: 
10348diff
changeset | 301 | (*Matching of cterms*) | 
| 16656 | 302 | fun gen_cterm_match match | 
| 20512 | 303 |     (ct1 as Cterm {t = t1, sorts = sorts1, ...},
 | 
| 20815 | 304 |      ct2 as Cterm {t = t2, sorts = sorts2, maxidx = maxidx2, ...}) =
 | 
| 10416 
5b33e732e459
- Moved rewriting functions to meta_simplifier.ML
 berghofe parents: 
10348diff
changeset | 305 | let | 
| 16656 | 306 | val thy_ref = merge_thys0 ct1 ct2; | 
| 18184 | 307 | val (Tinsts, tinsts) = match (Theory.deref thy_ref) (t1, t2) (Vartab.empty, Vartab.empty); | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 308 | val sorts = Sorts.union sorts1 sorts2; | 
| 20512 | 309 | fun mk_cTinst ((a, i), (S, T)) = | 
| 310 |       (Ctyp {T = TVar ((a, i), S), thy_ref = thy_ref, maxidx = i, sorts = sorts},
 | |
| 20815 | 311 |        Ctyp {T = T, thy_ref = thy_ref, maxidx = maxidx2, sorts = sorts});
 | 
| 20512 | 312 | fun mk_ctinst ((x, i), (T, t)) = | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 313 | let val T = Envir.typ_subst_TVars Tinsts T in | 
| 20512 | 314 |         (Cterm {t = Var ((x, i), T), T = T, thy_ref = thy_ref, maxidx = i, sorts = sorts},
 | 
| 20815 | 315 |          Cterm {t = t, T = T, thy_ref = thy_ref, maxidx = maxidx2, sorts = sorts})
 | 
| 10416 
5b33e732e459
- Moved rewriting functions to meta_simplifier.ML
 berghofe parents: 
10348diff
changeset | 316 | end; | 
| 16656 | 317 | 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 | 318 | |
| 
5b33e732e459
- Moved rewriting functions to meta_simplifier.ML
 berghofe parents: 
10348diff
changeset | 319 | val cterm_match = gen_cterm_match Pattern.match; | 
| 
5b33e732e459
- Moved rewriting functions to meta_simplifier.ML
 berghofe parents: 
10348diff
changeset | 320 | val cterm_first_order_match = gen_cterm_match Pattern.first_order_match; | 
| 
5b33e732e459
- Moved rewriting functions to meta_simplifier.ML
 berghofe parents: 
10348diff
changeset | 321 | |
| 
5b33e732e459
- Moved rewriting functions to meta_simplifier.ML
 berghofe parents: 
10348diff
changeset | 322 | (*Incrementing indexes*) | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 323 | fun cterm_incr_indexes i (ct as Cterm {thy_ref, t, T, maxidx, sorts}) =
 | 
| 22584 | 324 |   if i < 0 then raise CTERM ("negative increment", [ct])
 | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 325 | else if i = 0 then ct | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 326 |   else Cterm {thy_ref = thy_ref, t = Logic.incr_indexes ([], i) t,
 | 
| 16884 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 wenzelm parents: 
16847diff
changeset | 327 | T = Logic.incr_tvar i T, maxidx = maxidx + i, sorts = sorts}; | 
| 10416 
5b33e732e459
- Moved rewriting functions to meta_simplifier.ML
 berghofe parents: 
10348diff
changeset | 328 | |
| 2509 | 329 | |
| 330 | ||
| 387 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 331 | (*** Meta theorems ***) | 
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
225diff
changeset | 332 | |
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 333 | type tag = string * string list; | 
| 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 334 | |
| 22237 
bb9b1c8a8a95
old-fashioned abstype ctyp/cterm/thm prevents accidental equality tests;
 wenzelm parents: 
21975diff
changeset | 335 | abstype thm = Thm of | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 336 |  {thy_ref: theory_ref,         (*dynamic reference to theory*)
 | 
| 11518 | 337 | der: bool * Pt.proof, (*derivation*) | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 338 | tags: tag list, (*additional annotations/comments*) | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 339 | maxidx: int, (*maximum index of any Var or TVar*) | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 340 | shyps: sort list, (*sort hypotheses as ordered list*) | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 341 | hyps: term list, (*hypotheses as ordered list*) | 
| 13658 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 berghofe parents: 
13642diff
changeset | 342 | tpairs: (term * term) list, (*flex-flex pairs*) | 
| 22237 
bb9b1c8a8a95
old-fashioned abstype ctyp/cterm/thm prevents accidental equality tests;
 wenzelm parents: 
21975diff
changeset | 343 | prop: term} (*conclusion*) | 
| 
bb9b1c8a8a95
old-fashioned abstype ctyp/cterm/thm prevents accidental equality tests;
 wenzelm parents: 
21975diff
changeset | 344 | with | 
| 0 | 345 | |
| 22365 | 346 | (*attributes subsume any kind of rules or context modifiers*) | 
| 347 | type attribute = Context.generic * thm -> Context.generic * thm; | |
| 348 | ||
| 16725 | 349 | (*errors involving theorems*) | 
| 350 | exception THM of string * int * thm list; | |
| 13658 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 berghofe parents: 
13642diff
changeset | 351 | |
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 352 | fun rep_thm (Thm {thy_ref, der, tags, maxidx, shyps, hyps, tpairs, prop}) =
 | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 353 | let val thy = Theory.deref thy_ref in | 
| 22596 | 354 |    {thy = thy, der = der, tags = tags, maxidx = maxidx,
 | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 355 | shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop} | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 356 | end; | 
| 0 | 357 | |
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 358 | (*version of rep_thm returning cterms instead of terms*) | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 359 | fun crep_thm (Thm {thy_ref, der, tags, maxidx, shyps, hyps, tpairs, prop}) =
 | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 360 | let | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 361 | val thy = Theory.deref thy_ref; | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 362 |     fun cterm max t = Cterm {thy_ref = thy_ref, t = t, T = propT, maxidx = max, sorts = shyps};
 | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 363 | in | 
| 22596 | 364 |    {thy = thy, der = der, tags = tags, maxidx = maxidx, shyps = shyps,
 | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 365 | hyps = map (cterm ~1) hyps, | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 366 | tpairs = map (pairself (cterm maxidx)) tpairs, | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 367 | prop = cterm maxidx prop} | 
| 1517 | 368 | end; | 
| 369 | ||
| 16725 | 370 | fun terms_of_tpairs tpairs = fold_rev (fn (t, u) => cons t o cons u) tpairs []; | 
| 371 | ||
| 372 | fun eq_tpairs ((t, u), (t', u')) = t aconv t' andalso u aconv u'; | |
| 18944 | 373 | 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 | 374 | val maxidx_tpairs = fold (fn (t, u) => Term.maxidx_term t #> Term.maxidx_term u); | 
| 16725 | 375 | |
| 376 | fun attach_tpairs tpairs prop = | |
| 377 | Logic.list_implies (map Logic.mk_equals tpairs, prop); | |
| 378 | ||
| 379 | fun full_prop_of (Thm {tpairs, prop, ...}) = attach_tpairs tpairs prop;
 | |
| 16945 | 380 | |
| 22365 | 381 | val union_hyps = OrdList.union Term.fast_term_ord; | 
| 382 | ||
| 16945 | 383 | |
| 384 | (* merge theories of cterms/thms; raise exception if incompatible *) | |
| 385 | ||
| 386 | fun merge_thys1 (Cterm {thy_ref = r1, ...}) (th as Thm {thy_ref = r2, ...}) =
 | |
| 387 | Theory.merge_refs (r1, r2) handle TERM (msg, _) => raise THM (msg, 0, [th]); | |
| 388 | ||
| 389 | fun merge_thys2 (th1 as Thm {thy_ref = r1, ...}) (th2 as Thm {thy_ref = r2, ...}) =
 | |
| 390 | Theory.merge_refs (r1, r2) handle TERM (msg, _) => raise THM (msg, 0, [th1, th2]); | |
| 391 | ||
| 0 | 392 | |
| 22365 | 393 | (* basic components *) | 
| 16135 | 394 | |
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 395 | fun theory_of_thm (Thm {thy_ref, ...}) = Theory.deref thy_ref;
 | 
| 19429 | 396 | fun maxidx_of (Thm {maxidx, ...}) = maxidx;
 | 
| 19910 | 397 | fun maxidx_thm th i = Int.max (maxidx_of th, i); | 
| 19881 | 398 | fun hyps_of (Thm {hyps, ...}) = hyps;
 | 
| 12803 | 399 | fun prop_of (Thm {prop, ...}) = prop;
 | 
| 13528 | 400 | fun proof_of (Thm {der = (_, proof), ...}) = proof;
 | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 401 | fun tpairs_of (Thm {tpairs, ...}) = tpairs;
 | 
| 0 | 402 | |
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 403 | val concl_of = Logic.strip_imp_concl o prop_of; | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 404 | val prems_of = Logic.strip_imp_prems o prop_of; | 
| 21576 | 405 | val nprems_of = Logic.count_prems o prop_of; | 
| 19305 | 406 | fun no_prems th = nprems_of th = 0; | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 407 | |
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 408 | fun major_prem_of th = | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 409 | (case prems_of th of | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 410 | prem :: _ => Logic.strip_assums_concl prem | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 411 |   | [] => raise THM ("major_prem_of: rule with no premises", 0, [th]));
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 412 | |
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 413 | (*the statement of any thm is a cterm*) | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 414 | fun cprop_of (Thm {thy_ref, maxidx, shyps, prop, ...}) =
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 415 |   Cterm {thy_ref = thy_ref, maxidx = maxidx, T = propT, t = prop, sorts = shyps};
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 416 | |
| 18145 | 417 | fun cprem_of (th as Thm {thy_ref, maxidx, shyps, prop, ...}) i =
 | 
| 18035 | 418 |   Cterm {thy_ref = thy_ref, maxidx = maxidx, T = propT, sorts = shyps,
 | 
| 18145 | 419 |     t = Logic.nth_prem (i, prop) handle TERM _ => raise THM ("cprem_of", i, [th])};
 | 
| 18035 | 420 | |
| 16656 | 421 | (*explicit transfer to a super theory*) | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 422 | fun transfer thy' thm = | 
| 3895 | 423 | let | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 424 |     val Thm {thy_ref, der, tags, maxidx, shyps, hyps, tpairs, prop} = thm;
 | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 425 | val thy = Theory.deref thy_ref; | 
| 3895 | 426 | in | 
| 16945 | 427 | if not (subthy (thy, thy')) then | 
| 428 |       raise THM ("transfer: not a super theory", 0, [thm])
 | |
| 429 | else if eq_thy (thy, thy') then thm | |
| 430 | else | |
| 431 |       Thm {thy_ref = Theory.self_ref thy',
 | |
| 432 | der = der, | |
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 433 | tags = tags, | 
| 16945 | 434 | maxidx = maxidx, | 
| 435 | shyps = shyps, | |
| 436 | hyps = hyps, | |
| 437 | tpairs = tpairs, | |
| 438 | prop = prop} | |
| 3895 | 439 | end; | 
| 387 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 440 | |
| 16945 | 441 | (*explicit weakening: maps |- B to A |- B*) | 
| 442 | fun weaken raw_ct th = | |
| 443 | let | |
| 20261 | 444 |     val ct as Cterm {t = A, T, sorts, maxidx = maxidxA, ...} = adjust_maxidx_cterm ~1 raw_ct;
 | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 445 |     val Thm {der, tags, maxidx, shyps, hyps, tpairs, prop, ...} = th;
 | 
| 16945 | 446 | in | 
| 447 | if T <> propT then | |
| 448 |       raise THM ("weaken: assumptions must have type prop", 0, [])
 | |
| 449 | else if maxidxA <> ~1 then | |
| 450 |       raise THM ("weaken: assumptions may not contain schematic variables", maxidxA, [])
 | |
| 451 | else | |
| 452 |       Thm {thy_ref = merge_thys1 ct th,
 | |
| 453 | der = der, | |
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 454 | tags = tags, | 
| 16945 | 455 | maxidx = maxidx, | 
| 456 | shyps = Sorts.union sorts shyps, | |
| 22365 | 457 | hyps = OrdList.insert Term.fast_term_ord A hyps, | 
| 16945 | 458 | tpairs = tpairs, | 
| 459 | prop = prop} | |
| 460 | end; | |
| 16656 | 461 | |
| 462 | ||
| 0 | 463 | |
| 1238 | 464 | (** sort contexts of theorems **) | 
| 465 | ||
| 16656 | 466 | fun present_sorts (Thm {hyps, tpairs, prop, ...}) =
 | 
| 467 | fold (fn (t, u) => Sorts.insert_term t o Sorts.insert_term u) tpairs | |
| 468 | (Sorts.insert_terms hyps (Sorts.insert_term prop [])); | |
| 1238 | 469 | |
| 7642 | 470 | (*remove extra sorts that are non-empty by virtue of type signature information*) | 
| 471 | fun strip_shyps (thm as Thm {shyps = [], ...}) = thm
 | |
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 472 |   | strip_shyps (thm as Thm {thy_ref, der, tags, maxidx, shyps, hyps, tpairs, prop}) =
 | 
| 7642 | 473 | let | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 474 | val thy = Theory.deref thy_ref; | 
| 16656 | 475 | val shyps' = | 
| 476 | if Sign.all_sorts_nonempty thy then [] | |
| 477 | else | |
| 478 | let | |
| 479 | val present = present_sorts thm; | |
| 480 | val extra = Sorts.subtract present shyps; | |
| 481 | val witnessed = map #2 (Sign.witness_sorts thy present extra); | |
| 482 | in Sorts.subtract witnessed shyps end; | |
| 7642 | 483 | in | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 484 |         Thm {thy_ref = thy_ref, der = der, tags = tags, maxidx = maxidx,
 | 
| 16656 | 485 | shyps = shyps', hyps = hyps, tpairs = tpairs, prop = prop} | 
| 7642 | 486 | end; | 
| 1238 | 487 | |
| 16656 | 488 | (*dangling sort constraints of a thm*) | 
| 489 | fun extra_shyps (th as Thm {shyps, ...}) = Sorts.subtract (present_sorts th) shyps;
 | |
| 490 | ||
| 1238 | 491 | |
| 492 | ||
| 1529 | 493 | (** Axioms **) | 
| 387 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 494 | |
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 495 | (*look up the named axiom in the theory or its ancestors*) | 
| 15672 | 496 | fun get_axiom_i theory name = | 
| 387 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 497 | let | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 498 | fun get_ax thy = | 
| 22685 | 499 | Symtab.lookup (Theory.axiom_table thy) name | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 500 | |> Option.map (fn prop => | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 501 |           Thm {thy_ref = Theory.self_ref thy,
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 502 | der = Pt.infer_derivs' I (false, Pt.axm_proof name prop), | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 503 | tags = [], | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 504 | maxidx = maxidx_of_term prop, | 
| 16656 | 505 | shyps = may_insert_term_sorts thy prop [], | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 506 | hyps = [], | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 507 | tpairs = [], | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 508 | prop = prop}); | 
| 387 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 509 | in | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 510 | (case get_first get_ax (theory :: Theory.ancestors_of theory) of | 
| 15531 | 511 | SOME thm => thm | 
| 512 |     | NONE => raise THEORY ("No axiom " ^ quote name, [theory]))
 | |
| 387 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 513 | end; | 
| 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 514 | |
| 16352 | 515 | fun get_axiom thy = | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 516 | get_axiom_i thy o NameSpace.intern (Theory.axiom_space thy); | 
| 15672 | 517 | |
| 20884 | 518 | fun def_name c = c ^ "_def"; | 
| 519 | ||
| 520 | fun def_name_optional c "" = def_name c | |
| 521 | | def_name_optional _ name = name; | |
| 522 | ||
| 6368 | 523 | fun get_def thy = get_axiom thy o def_name; | 
| 4847 | 524 | |
| 1529 | 525 | |
| 776 
df8f91c0e57c
improved axioms_of: returns thms as the manual says;
 wenzelm parents: 
721diff
changeset | 526 | (*return additional axioms of this theory node*) | 
| 
df8f91c0e57c
improved axioms_of: returns thms as the manual says;
 wenzelm parents: 
721diff
changeset | 527 | fun axioms_of thy = | 
| 22685 | 528 | map (fn s => (s, get_axiom_i thy s)) (Symtab.keys (Theory.axiom_table thy)); | 
| 776 
df8f91c0e57c
improved axioms_of: returns thms as the manual says;
 wenzelm parents: 
721diff
changeset | 529 | |
| 6089 | 530 | |
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 531 | (* official name and additional tags *) | 
| 6089 | 532 | |
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 533 | fun get_name (Thm {hyps, prop, der = (_, prf), ...}) =
 | 
| 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 534 | Pt.get_name hyps prop prf; | 
| 4018 | 535 | |
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 536 | fun put_name name (Thm {thy_ref, der = (ora, prf), tags, maxidx, shyps, hyps, tpairs = [], prop}) =
 | 
| 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 537 |       Thm {thy_ref = thy_ref,
 | 
| 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 538 | der = (ora, Pt.thm_proof (Theory.deref thy_ref) name hyps prop prf), | 
| 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 539 | tags = tags, maxidx = maxidx, shyps = shyps, hyps = hyps, tpairs = [], prop = prop} | 
| 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 540 |   | put_name _ thm = raise THM ("name_thm: unsolved flex-flex constraints", 0, [thm]);
 | 
| 6089 | 541 | |
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 542 | val get_tags = #tags o rep_thm; | 
| 6089 | 543 | |
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 544 | fun map_tags f (Thm {thy_ref, der, tags, maxidx, shyps, hyps, tpairs, prop}) =
 | 
| 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 545 |   Thm {thy_ref = thy_ref, der = der, tags = f tags, maxidx = maxidx,
 | 
| 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 546 | shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop}; | 
| 0 | 547 | |
| 548 | ||
| 1529 | 549 | (*Compression of theorems -- a separate rule, not integrated with the others, | 
| 550 | as it could be slow.*) | |
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 551 | fun compress (Thm {thy_ref, der, tags, maxidx, shyps, hyps, tpairs, prop}) =
 | 
| 16991 | 552 | let val thy = Theory.deref thy_ref in | 
| 553 |     Thm {thy_ref = thy_ref,
 | |
| 554 | der = der, | |
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 555 | tags = tags, | 
| 16991 | 556 | maxidx = maxidx, | 
| 557 | shyps = shyps, | |
| 558 | hyps = map (Compress.term thy) hyps, | |
| 559 | tpairs = map (pairself (Compress.term thy)) tpairs, | |
| 560 | prop = Compress.term thy prop} | |
| 561 | end; | |
| 16945 | 562 | |
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 563 | fun adjust_maxidx_thm i (th as Thm {thy_ref, der, tags, maxidx, shyps, hyps, tpairs, prop}) =
 | 
| 20261 | 564 | if maxidx = i then th | 
| 565 | else if maxidx < i then | |
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 566 |     Thm {maxidx = i, thy_ref = thy_ref, der = der, tags = tags, shyps = shyps,
 | 
| 20261 | 567 | hyps = hyps, tpairs = tpairs, prop = prop} | 
| 568 | else | |
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 569 |     Thm {maxidx = Int.max (maxidx_tpairs tpairs (maxidx_of_term prop), i), thy_ref = thy_ref,
 | 
| 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 570 | der = der, tags = tags, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop}; | 
| 564 | 571 | |
| 387 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 572 | |
| 2509 | 573 | |
| 1529 | 574 | (*** Meta rules ***) | 
| 0 | 575 | |
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 576 | (** primitive rules **) | 
| 0 | 577 | |
| 16656 | 578 | (*The assumption rule A |- A*) | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 579 | fun assume raw_ct = | 
| 20261 | 580 |   let val Cterm {thy_ref, t = prop, T, maxidx, sorts} = adjust_maxidx_cterm ~1 raw_ct in
 | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 581 | if T <> propT then | 
| 19230 | 582 |       raise THM ("assume: prop", 0, [])
 | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 583 | else if maxidx <> ~1 then | 
| 19230 | 584 |       raise THM ("assume: variables", maxidx, [])
 | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 585 |     else Thm {thy_ref = thy_ref,
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 586 | der = Pt.infer_derivs' I (false, Pt.Hyp prop), | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 587 | tags = [], | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 588 | maxidx = ~1, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 589 | shyps = sorts, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 590 | hyps = [prop], | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 591 | tpairs = [], | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 592 | prop = prop} | 
| 0 | 593 | end; | 
| 594 | ||
| 1220 | 595 | (*Implication introduction | 
| 3529 | 596 | [A] | 
| 597 | : | |
| 598 | B | |
| 1220 | 599 | ------- | 
| 600 | A ==> B | |
| 601 | *) | |
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 602 | fun implies_intr | 
| 16679 | 603 |     (ct as Cterm {t = A, T, maxidx = maxidxA, sorts, ...})
 | 
| 604 |     (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 | 605 | if T <> propT then | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 606 |     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 | 607 | else | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 608 |     Thm {thy_ref = merge_thys1 ct th,
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 609 | der = Pt.infer_derivs' (Pt.implies_intr_proof A) der, | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 610 | tags = [], | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 611 | maxidx = Int.max (maxidxA, maxidx), | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 612 | shyps = Sorts.union sorts shyps, | 
| 22365 | 613 | hyps = OrdList.remove Term.fast_term_ord A hyps, | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 614 | tpairs = tpairs, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 615 | prop = implies $ A $ prop}; | 
| 0 | 616 | |
| 1529 | 617 | |
| 1220 | 618 | (*Implication elimination | 
| 619 | A ==> B A | |
| 620 | ------------ | |
| 621 | B | |
| 622 | *) | |
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 623 | fun implies_elim thAB thA = | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 624 | let | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 625 |     val Thm {maxidx = maxA, der = derA, hyps = hypsA, shyps = shypsA, tpairs = tpairsA,
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 626 | prop = propA, ...} = thA | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 627 |     and Thm {der, maxidx, hyps, shyps, tpairs, prop, ...} = thAB;
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 628 |     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 | 629 | in | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 630 | case prop of | 
| 20512 | 631 |       Const ("==>", _) $ A $ B =>
 | 
| 632 | if A aconv propA then | |
| 16656 | 633 |           Thm {thy_ref = merge_thys2 thAB thA,
 | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 634 | der = Pt.infer_derivs (curry Pt.%%) der derA, | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 635 | tags = [], | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 636 | maxidx = Int.max (maxA, maxidx), | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 637 | shyps = Sorts.union shypsA shyps, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 638 | hyps = union_hyps hypsA hyps, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 639 | tpairs = union_tpairs tpairsA tpairs, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 640 | prop = B} | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 641 | else err () | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 642 | | _ => err () | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 643 | end; | 
| 250 | 644 | |
| 1220 | 645 | (*Forall introduction. The Free or Var x must not be free in the hypotheses. | 
| 16656 | 646 | [x] | 
| 647 | : | |
| 648 | A | |
| 649 | ------ | |
| 650 | !!x. A | |
| 1220 | 651 | *) | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 652 | fun forall_intr | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 653 |     (ct as Cterm {t = x, T, sorts, ...})
 | 
| 16679 | 654 |     (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 | 655 | let | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 656 | fun result a = | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 657 |       Thm {thy_ref = merge_thys1 ct th,
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 658 | der = Pt.infer_derivs' (Pt.forall_intr_proof x a) der, | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 659 | tags = [], | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 660 | maxidx = maxidx, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 661 | shyps = Sorts.union sorts shyps, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 662 | hyps = hyps, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 663 | tpairs = tpairs, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 664 | prop = all T $ Abs (a, T, abstract_over (x, prop))}; | 
| 21798 | 665 | fun check_occs a x ts = | 
| 16847 
8fc160b12e73
invoke_oracle: do not keep theory value, but theory_ref;
 wenzelm parents: 
16725diff
changeset | 666 | if exists (fn t => Logic.occs (x, t)) ts then | 
| 21798 | 667 |         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 | 668 | else (); | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 669 | in | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 670 | case x of | 
| 21798 | 671 | Free (a, _) => (check_occs a x hyps; check_occs a x (terms_of_tpairs tpairs); result a) | 
| 672 | | Var ((a, _), _) => (check_occs a x (terms_of_tpairs tpairs); result a) | |
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 673 |     | _ => raise THM ("forall_intr: not a variable", 0, [th])
 | 
| 0 | 674 | end; | 
| 675 | ||
| 1220 | 676 | (*Forall elimination | 
| 16656 | 677 | !!x. A | 
| 1220 | 678 | ------ | 
| 679 | A[t/x] | |
| 680 | *) | |
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 681 | fun forall_elim | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 682 |     (ct as Cterm {t, T, maxidx = maxt, sorts, ...})
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 683 |     (th as Thm {der, maxidx, shyps, hyps, tpairs, prop, ...}) =
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 684 | (case prop of | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 685 |     Const ("all", Type ("fun", [Type ("fun", [qary, _]), _])) $ A =>
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 686 | if T <> qary then | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 687 |         raise THM ("forall_elim: type mismatch", 0, [th])
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 688 | else | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 689 |         Thm {thy_ref = merge_thys1 ct th,
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 690 | der = Pt.infer_derivs' (Pt.% o rpair (SOME t)) der, | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 691 | tags = [], | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 692 | maxidx = Int.max (maxidx, maxt), | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 693 | shyps = Sorts.union sorts shyps, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 694 | hyps = hyps, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 695 | tpairs = tpairs, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 696 | prop = Term.betapply (A, t)} | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 697 |   | _ => raise THM ("forall_elim: not quantified", 0, [th]));
 | 
| 0 | 698 | |
| 699 | ||
| 1220 | 700 | (* Equality *) | 
| 0 | 701 | |
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 702 | (*Reflexivity | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 703 | t == t | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 704 | *) | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 705 | fun reflexive (ct as Cterm {thy_ref, t, T, maxidx, sorts}) =
 | 
| 16656 | 706 |   Thm {thy_ref = thy_ref,
 | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 707 | der = Pt.infer_derivs' I (false, Pt.reflexive), | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 708 | tags = [], | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 709 | maxidx = maxidx, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 710 | shyps = sorts, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 711 | hyps = [], | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 712 | tpairs = [], | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 713 | prop = Logic.mk_equals (t, t)}; | 
| 0 | 714 | |
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 715 | (*Symmetry | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 716 | t == u | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 717 | ------ | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 718 | u == t | 
| 1220 | 719 | *) | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 720 | fun symmetric (th as Thm {thy_ref, der, maxidx, shyps, hyps, tpairs, prop, ...}) =
 | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 721 | (case prop of | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 722 |     (eq as Const ("==", Type (_, [T, _]))) $ t $ u =>
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 723 |       Thm {thy_ref = thy_ref,
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 724 | der = Pt.infer_derivs' Pt.symmetric der, | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 725 | tags = [], | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 726 | maxidx = maxidx, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 727 | shyps = shyps, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 728 | hyps = hyps, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 729 | tpairs = tpairs, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 730 | prop = eq $ u $ t} | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 731 |     | _ => raise THM ("symmetric", 0, [th]));
 | 
| 0 | 732 | |
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 733 | (*Transitivity | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 734 | t1 == u u == t2 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 735 | ------------------ | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 736 | t1 == t2 | 
| 1220 | 737 | *) | 
| 0 | 738 | fun transitive th1 th2 = | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 739 | let | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 740 |     val Thm {der = der1, maxidx = max1, hyps = hyps1, shyps = shyps1, tpairs = tpairs1,
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 741 | prop = prop1, ...} = th1 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 742 |     and Thm {der = der2, maxidx = max2, hyps = hyps2, shyps = shyps2, tpairs = tpairs2,
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 743 | prop = prop2, ...} = th2; | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 744 |     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 | 745 | in | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 746 | case (prop1, prop2) of | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 747 |       ((eq as Const ("==", Type (_, [T, _]))) $ t1 $ u, Const ("==", _) $ u' $ t2) =>
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 748 | if not (u aconv u') then err "middle term" | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 749 | else | 
| 16656 | 750 |           Thm {thy_ref = merge_thys2 th1 th2,
 | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 751 | der = Pt.infer_derivs (Pt.transitive u T) der1 der2, | 
| 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 = Int.max (max1, max2), | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 754 | shyps = Sorts.union shyps1 shyps2, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 755 | hyps = union_hyps hyps1 hyps2, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 756 | tpairs = union_tpairs tpairs1 tpairs2, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 757 | prop = eq $ t1 $ t2} | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 758 | | _ => err "premises" | 
| 0 | 759 | end; | 
| 760 | ||
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 761 | (*Beta-conversion | 
| 16656 | 762 | (%x. t)(u) == t[u/x] | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 763 | fully beta-reduces the term if full = true | 
| 10416 
5b33e732e459
- Moved rewriting functions to meta_simplifier.ML
 berghofe parents: 
10348diff
changeset | 764 | *) | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 765 | fun beta_conversion full (Cterm {thy_ref, t, T, maxidx, sorts}) =
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 766 | let val t' = | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 767 | if full then Envir.beta_norm t | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 768 | else | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 769 | (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 | 770 |       | _ => raise THM ("beta_conversion: not a redex", 0, []));
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 771 | in | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 772 |     Thm {thy_ref = thy_ref,
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 773 | der = Pt.infer_derivs' I (false, Pt.reflexive), | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 774 | tags = [], | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 775 | maxidx = maxidx, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 776 | shyps = sorts, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 777 | hyps = [], | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 778 | tpairs = [], | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 779 | prop = Logic.mk_equals (t, t')} | 
| 10416 
5b33e732e459
- Moved rewriting functions to meta_simplifier.ML
 berghofe parents: 
10348diff
changeset | 780 | end; | 
| 
5b33e732e459
- Moved rewriting functions to meta_simplifier.ML
 berghofe parents: 
10348diff
changeset | 781 | |
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 782 | fun eta_conversion (Cterm {thy_ref, t, T, maxidx, sorts}) =
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 783 |   Thm {thy_ref = thy_ref,
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 784 | der = Pt.infer_derivs' I (false, Pt.reflexive), | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 785 | tags = [], | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 786 | maxidx = maxidx, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 787 | shyps = sorts, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 788 | hyps = [], | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 789 | tpairs = [], | 
| 18944 | 790 | prop = Logic.mk_equals (t, Envir.eta_contract t)}; | 
| 0 | 791 | |
| 792 | (*The abstraction rule. The Free or Var x must not be free in the hypotheses. | |
| 793 | 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 | 794 | t == u | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 795 | -------------- | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 796 | %x. t == %x. u | 
| 1220 | 797 | *) | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 798 | fun abstract_rule a | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 799 |     (Cterm {t = x, T, sorts, ...})
 | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 800 |     (th as Thm {thy_ref, der, maxidx, hyps, shyps, tpairs, prop, ...}) =
 | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 801 | let | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 802 | val (t, u) = Logic.dest_equals prop | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 803 |       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 | 804 | val result = | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 805 |       Thm {thy_ref = thy_ref,
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 806 | der = Pt.infer_derivs' (Pt.abstract_rule x a) der, | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 807 | tags = [], | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 808 | maxidx = maxidx, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 809 | shyps = Sorts.union sorts shyps, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 810 | hyps = hyps, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 811 | tpairs = tpairs, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 812 | prop = Logic.mk_equals | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 813 | (Abs (a, T, abstract_over (x, t)), Abs (a, T, abstract_over (x, u)))}; | 
| 21798 | 814 | fun check_occs a x ts = | 
| 16847 
8fc160b12e73
invoke_oracle: do not keep theory value, but theory_ref;
 wenzelm parents: 
16725diff
changeset | 815 | if exists (fn t => Logic.occs (x, t)) ts then | 
| 21798 | 816 |         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 | 817 | else (); | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 818 | in | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 819 | case x of | 
| 21798 | 820 | Free (a, _) => (check_occs a x hyps; check_occs a x (terms_of_tpairs tpairs); result) | 
| 821 | | Var ((a, _), _) => (check_occs a x (terms_of_tpairs tpairs); result) | |
| 822 |     | _ => raise THM ("abstract_rule: not a variable", 0, [th])
 | |
| 0 | 823 | end; | 
| 824 | ||
| 825 | (*The combination rule | |
| 3529 | 826 | f == g t == u | 
| 827 | -------------- | |
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 828 | f t == g u | 
| 1220 | 829 | *) | 
| 0 | 830 | fun combination th1 th2 = | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 831 | let | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 832 |     val Thm {der = der1, maxidx = max1, shyps = shyps1, hyps = hyps1, tpairs = tpairs1,
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 833 | prop = prop1, ...} = th1 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 834 |     and Thm {der = der2, maxidx = max2, shyps = shyps2, hyps = hyps2, tpairs = tpairs2,
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 835 | prop = prop2, ...} = th2; | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 836 | fun chktypes fT tT = | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 837 | (case fT of | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 838 |         Type ("fun", [T1, T2]) =>
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 839 | if T1 <> tT then | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 840 |             raise THM ("combination: types", 0, [th1, th2])
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 841 | else () | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 842 |       | _ => raise THM ("combination: not function type", 0, [th1, th2]));
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 843 | in | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 844 | case (prop1, prop2) of | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 845 |       (Const ("==", Type ("fun", [fT, _])) $ f $ g,
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 846 |        Const ("==", Type ("fun", [tT, _])) $ t $ u) =>
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 847 | (chktypes fT tT; | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 848 |           Thm {thy_ref = merge_thys2 th1 th2,
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 849 | der = Pt.infer_derivs (Pt.combination f g t u fT) der1 der2, | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 850 | tags = [], | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 851 | maxidx = Int.max (max1, max2), | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 852 | shyps = Sorts.union shyps1 shyps2, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 853 | hyps = union_hyps hyps1 hyps2, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 854 | tpairs = union_tpairs tpairs1 tpairs2, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 855 | prop = Logic.mk_equals (f $ t, g $ u)}) | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 856 |      | _ => raise THM ("combination: premises", 0, [th1, th2])
 | 
| 0 | 857 | end; | 
| 858 | ||
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 859 | (*Equality introduction | 
| 3529 | 860 | A ==> B B ==> A | 
| 861 | ---------------- | |
| 862 | A == B | |
| 1220 | 863 | *) | 
| 0 | 864 | fun equal_intr th1 th2 = | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 865 | let | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 866 |     val Thm {der = der1, maxidx = max1, shyps = shyps1, hyps = hyps1, tpairs = tpairs1,
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 867 | prop = prop1, ...} = th1 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 868 |     and Thm {der = der2, maxidx = max2, shyps = shyps2, hyps = hyps2, tpairs = tpairs2,
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 869 | prop = prop2, ...} = th2; | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 870 |     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 | 871 | in | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 872 | case (prop1, prop2) of | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 873 |       (Const("==>", _) $ A $ B, Const("==>", _) $ B' $ A') =>
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 874 | if A aconv A' andalso B aconv B' then | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 875 |           Thm {thy_ref = merge_thys2 th1 th2,
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 876 | der = Pt.infer_derivs (Pt.equal_intr A B) der1 der2, | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 877 | tags = [], | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 878 | maxidx = Int.max (max1, max2), | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 879 | shyps = Sorts.union shyps1 shyps2, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 880 | hyps = union_hyps hyps1 hyps2, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 881 | tpairs = union_tpairs tpairs1 tpairs2, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 882 | prop = Logic.mk_equals (A, B)} | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 883 | else err "not equal" | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 884 | | _ => err "premises" | 
| 1529 | 885 | end; | 
| 886 | ||
| 887 | (*The equal propositions rule | |
| 3529 | 888 | A == B A | 
| 1529 | 889 | --------- | 
| 890 | B | |
| 891 | *) | |
| 892 | fun equal_elim th1 th2 = | |
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 893 | let | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 894 |     val Thm {der = der1, maxidx = max1, shyps = shyps1, hyps = hyps1,
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 895 | tpairs = tpairs1, prop = prop1, ...} = th1 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 896 |     and Thm {der = der2, maxidx = max2, shyps = shyps2, hyps = hyps2,
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 897 | tpairs = tpairs2, prop = prop2, ...} = th2; | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 898 |     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 | 899 | in | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 900 | case prop1 of | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 901 |       Const ("==", _) $ A $ B =>
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 902 | if prop2 aconv A then | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 903 |           Thm {thy_ref = merge_thys2 th1 th2,
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 904 | der = Pt.infer_derivs (Pt.equal_elim A B) der1 der2, | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 905 | tags = [], | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 906 | maxidx = Int.max (max1, max2), | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 907 | shyps = Sorts.union shyps1 shyps2, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 908 | hyps = union_hyps hyps1 hyps2, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 909 | tpairs = union_tpairs tpairs1 tpairs2, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 910 | prop = B} | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 911 | else err "not equal" | 
| 1529 | 912 | | _ => err"major premise" | 
| 913 | end; | |
| 0 | 914 | |
| 1220 | 915 | |
| 916 | ||
| 0 | 917 | (**** Derived rules ****) | 
| 918 | ||
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 919 | (*Smash unifies the list of term pairs leaving no flex-flex pairs. | 
| 250 | 920 | Instantiates the theorem and deletes trivial tpairs. | 
| 0 | 921 | Resulting sequence may contain multiple elements if the tpairs are | 
| 922 | not all flex-flex. *) | |
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 923 | fun flexflex_rule (th as Thm {thy_ref, der, maxidx, shyps, hyps, tpairs, prop, ...}) =
 | 
| 19861 | 924 | Unify.smash_unifiers (Theory.deref thy_ref) tpairs (Envir.empty maxidx) | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 925 | |> Seq.map (fn env => | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 926 | if Envir.is_empty env then th | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 927 | else | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 928 | let | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 929 | val tpairs' = tpairs |> map (pairself (Envir.norm_term env)) | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 930 | (*remove trivial tpairs, of the form t==t*) | 
| 16884 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 wenzelm parents: 
16847diff
changeset | 931 | |> filter_out (op aconv); | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 932 | val prop' = Envir.norm_term env prop; | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 933 | in | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 934 |           Thm {thy_ref = thy_ref,
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 935 | der = Pt.infer_derivs' (Pt.norm_proof' env) der, | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 936 | tags = [], | 
| 16711 | 937 | maxidx = maxidx_tpairs tpairs' (maxidx_of_term prop'), | 
| 16656 | 938 | shyps = may_insert_env_sorts (Theory.deref thy_ref) env shyps, | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 939 | hyps = hyps, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 940 | tpairs = tpairs', | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 941 | prop = prop'} | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 942 | end); | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 943 | |
| 0 | 944 | |
| 19910 | 945 | (*Generalization of fixed variables | 
| 946 | A | |
| 947 | -------------------- | |
| 948 | A[?'a/'a, ?x/x, ...] | |
| 949 | *) | |
| 950 | ||
| 951 | fun generalize ([], []) _ th = th | |
| 952 | | generalize (tfrees, frees) idx th = | |
| 953 | let | |
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 954 |         val Thm {thy_ref, der, maxidx, shyps, hyps, tpairs, prop, ...} = th;
 | 
| 19910 | 955 |         val _ = idx <= maxidx andalso raise THM ("generalize: bad index", idx, [th]);
 | 
| 956 | ||
| 957 | val bad_type = if null tfrees then K false else | |
| 958 | Term.exists_subtype (fn TFree (a, _) => member (op =) tfrees a | _ => false); | |
| 959 | fun bad_term (Free (x, T)) = bad_type T orelse member (op =) frees x | |
| 960 | | bad_term (Var (_, T)) = bad_type T | |
| 961 | | bad_term (Const (_, T)) = bad_type T | |
| 962 | | bad_term (Abs (_, T, t)) = bad_type T orelse bad_term t | |
| 963 | | bad_term (t $ u) = bad_term t orelse bad_term u | |
| 964 | | bad_term (Bound _) = false; | |
| 965 | val _ = exists bad_term hyps andalso | |
| 966 |           raise THM ("generalize: variable free in assumptions", 0, [th]);
 | |
| 967 | ||
| 20512 | 968 | val gen = TermSubst.generalize (tfrees, frees) idx; | 
| 19910 | 969 | val prop' = gen prop; | 
| 970 | val tpairs' = map (pairself gen) tpairs; | |
| 971 | val maxidx' = maxidx_tpairs tpairs' (maxidx_of_term prop'); | |
| 972 | in | |
| 973 |         Thm {
 | |
| 974 | thy_ref = thy_ref, | |
| 975 | der = Pt.infer_derivs' (Pt.generalize (tfrees, frees) idx) der, | |
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 976 | tags = [], | 
| 19910 | 977 | maxidx = maxidx', | 
| 978 | shyps = shyps, | |
| 979 | hyps = hyps, | |
| 980 | tpairs = tpairs', | |
| 981 | prop = prop'} | |
| 982 | end; | |
| 983 | ||
| 984 | ||
| 22584 | 985 | (*Instantiation of schematic variables | 
| 16656 | 986 | A | 
| 987 | -------------------- | |
| 988 | A[t1/v1, ..., tn/vn] | |
| 1220 | 989 | *) | 
| 0 | 990 | |
| 6928 | 991 | local | 
| 992 | ||
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 993 | fun pretty_typing thy t T = | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 994 | Pretty.block [Sign.pretty_term thy t, Pretty.str " ::", Pretty.brk 1, Sign.pretty_typ thy T]; | 
| 15797 | 995 | |
| 16884 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 wenzelm parents: 
16847diff
changeset | 996 | fun add_inst (ct, cu) (thy_ref, sorts) = | 
| 6928 | 997 | let | 
| 16884 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 wenzelm parents: 
16847diff
changeset | 998 |     val Cterm {t = t, T = T, ...} = ct
 | 
| 20512 | 999 |     and Cterm {t = u, T = U, sorts = sorts_u, maxidx = maxidx_u, ...} = cu;
 | 
| 16884 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 wenzelm parents: 
16847diff
changeset | 1000 | val thy_ref' = Theory.merge_refs (thy_ref, merge_thys0 ct cu); | 
| 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 wenzelm parents: 
16847diff
changeset | 1001 | val sorts' = Sorts.union sorts_u sorts; | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1002 | in | 
| 16884 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 wenzelm parents: 
16847diff
changeset | 1003 | (case t of Var v => | 
| 20512 | 1004 | if T = U then ((v, (u, maxidx_u)), (thy_ref', sorts')) | 
| 16884 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 wenzelm parents: 
16847diff
changeset | 1005 | else raise TYPE (Pretty.string_of (Pretty.block | 
| 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 wenzelm parents: 
16847diff
changeset | 1006 | [Pretty.str "instantiate: type conflict", | 
| 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 wenzelm parents: 
16847diff
changeset | 1007 | Pretty.fbrk, pretty_typing (Theory.deref thy_ref') t T, | 
| 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 wenzelm parents: 
16847diff
changeset | 1008 | Pretty.fbrk, pretty_typing (Theory.deref thy_ref') u U]), [T, U], [t, u]) | 
| 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 wenzelm parents: 
16847diff
changeset | 1009 | | _ => raise TYPE (Pretty.string_of (Pretty.block | 
| 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 wenzelm parents: 
16847diff
changeset | 1010 | [Pretty.str "instantiate: not a variable", | 
| 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 wenzelm parents: 
16847diff
changeset | 1011 | Pretty.fbrk, Sign.pretty_term (Theory.deref thy_ref') t]), [], [t])) | 
| 0 | 1012 | end; | 
| 1013 | ||
| 16884 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 wenzelm parents: 
16847diff
changeset | 1014 | fun add_instT (cT, cU) (thy_ref, sorts) = | 
| 16656 | 1015 | let | 
| 16884 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 wenzelm parents: 
16847diff
changeset | 1016 |     val Ctyp {T, thy_ref = thy_ref1, ...} = cT
 | 
| 20512 | 1017 |     and Ctyp {T = U, thy_ref = thy_ref2, sorts = sorts_U, maxidx = maxidx_U, ...} = cU;
 | 
| 16884 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 wenzelm parents: 
16847diff
changeset | 1018 | val thy_ref' = Theory.merge_refs (thy_ref, Theory.merge_refs (thy_ref1, thy_ref2)); | 
| 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 wenzelm parents: 
16847diff
changeset | 1019 | val thy' = Theory.deref thy_ref'; | 
| 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 wenzelm parents: 
16847diff
changeset | 1020 | val sorts' = Sorts.union sorts_U sorts; | 
| 16656 | 1021 | in | 
| 16884 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 wenzelm parents: 
16847diff
changeset | 1022 | (case T of TVar (v as (_, S)) => | 
| 20512 | 1023 | if Sign.of_sort thy' (U, S) then ((v, (U, maxidx_U)), (thy_ref', sorts')) | 
| 16656 | 1024 |       else raise TYPE ("Type not of sort " ^ Sign.string_of_sort thy' S, [U], [])
 | 
| 1025 | | _ => raise TYPE (Pretty.string_of (Pretty.block | |
| 15797 | 1026 | [Pretty.str "instantiate: not a type variable", | 
| 16656 | 1027 | Pretty.fbrk, Sign.pretty_typ thy' T]), [T], [])) | 
| 1028 | end; | |
| 0 | 1029 | |
| 6928 | 1030 | in | 
| 1031 | ||
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1032 | (*Left-to-right replacements: ctpairs = [..., (vi, ti), ...]. | 
| 0 | 1033 | Instantiates distinct Vars by terms of same type. | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1034 | Does NOT normalize the resulting theorem!*) | 
| 1529 | 1035 | fun instantiate ([], []) th = th | 
| 16884 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 wenzelm parents: 
16847diff
changeset | 1036 | | instantiate (instT, inst) th = | 
| 16656 | 1037 | let | 
| 16884 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 wenzelm parents: 
16847diff
changeset | 1038 |         val Thm {thy_ref, der, hyps, shyps, tpairs, prop, ...} = th;
 | 
| 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 wenzelm parents: 
16847diff
changeset | 1039 | val (inst', (instT', (thy_ref', shyps'))) = | 
| 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 wenzelm parents: 
16847diff
changeset | 1040 | (thy_ref, shyps) |> fold_map add_inst inst ||> fold_map add_instT instT; | 
| 20512 | 1041 | val subst = TermSubst.instantiate_maxidx (instT', inst'); | 
| 1042 | val (prop', maxidx1) = subst prop ~1; | |
| 1043 | val (tpairs', maxidx') = | |
| 1044 | fold_map (fn (t, u) => fn i => subst t i ||>> subst u) tpairs maxidx1; | |
| 16656 | 1045 | in | 
| 20545 
4c1068697159
instantiate: omit has_duplicates check, which is irrelevant for soundness;
 wenzelm parents: 
20512diff
changeset | 1046 |         Thm {thy_ref = thy_ref',
 | 
| 
4c1068697159
instantiate: omit has_duplicates check, which is irrelevant for soundness;
 wenzelm parents: 
20512diff
changeset | 1047 | der = Pt.infer_derivs' (fn d => | 
| 
4c1068697159
instantiate: omit has_duplicates check, which is irrelevant for soundness;
 wenzelm parents: 
20512diff
changeset | 1048 | Pt.instantiate (map (apsnd #1) instT', map (apsnd #1) inst') d) der, | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 1049 | tags = [], | 
| 20545 
4c1068697159
instantiate: omit has_duplicates check, which is irrelevant for soundness;
 wenzelm parents: 
20512diff
changeset | 1050 | maxidx = maxidx', | 
| 
4c1068697159
instantiate: omit has_duplicates check, which is irrelevant for soundness;
 wenzelm parents: 
20512diff
changeset | 1051 | shyps = shyps', | 
| 
4c1068697159
instantiate: omit has_duplicates check, which is irrelevant for soundness;
 wenzelm parents: 
20512diff
changeset | 1052 | hyps = hyps, | 
| 
4c1068697159
instantiate: omit has_duplicates check, which is irrelevant for soundness;
 wenzelm parents: 
20512diff
changeset | 1053 | tpairs = tpairs', | 
| 
4c1068697159
instantiate: omit has_duplicates check, which is irrelevant for soundness;
 wenzelm parents: 
20512diff
changeset | 1054 | prop = prop'} | 
| 16656 | 1055 | end | 
| 1056 | handle TYPE (msg, _, _) => raise THM (msg, 0, [th]); | |
| 6928 | 1057 | |
| 22584 | 1058 | fun instantiate_cterm ([], []) ct = ct | 
| 1059 | | instantiate_cterm (instT, inst) ct = | |
| 1060 | let | |
| 1061 |         val Cterm {thy_ref, t, T, sorts, ...} = ct;
 | |
| 1062 | val (inst', (instT', (thy_ref', sorts'))) = | |
| 1063 | (thy_ref, sorts) |> fold_map add_inst inst ||> fold_map add_instT instT; | |
| 1064 | val subst = TermSubst.instantiate_maxidx (instT', inst'); | |
| 1065 | val substT = TermSubst.instantiateT_maxidx instT'; | |
| 1066 | val (t', maxidx1) = subst t ~1; | |
| 1067 | val (T', maxidx') = substT T maxidx1; | |
| 1068 |       in Cterm {thy_ref = thy_ref', t = t', T = T', sorts = sorts', maxidx = maxidx'} end
 | |
| 1069 | handle TYPE (msg, _, _) => raise CTERM (msg, [ct]); | |
| 1070 | ||
| 6928 | 1071 | end; | 
| 1072 | ||
| 0 | 1073 | |
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1074 | (*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 | 1075 | A can contain Vars, not so for assume!*) | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1076 | fun trivial (Cterm {thy_ref, t =A, T, maxidx, sorts}) =
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1077 | if T <> propT then | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1078 |     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 | 1079 | else | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1080 |     Thm {thy_ref = thy_ref,
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1081 |       der = Pt.infer_derivs' I (false, Pt.AbsP ("H", NONE, Pt.PBound 0)),
 | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 1082 | tags = [], | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1083 | maxidx = maxidx, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1084 | shyps = sorts, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1085 | hyps = [], | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1086 | tpairs = [], | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1087 | prop = implies $ A $ A}; | 
| 0 | 1088 | |
| 1503 | 1089 | (*Axiom-scheme reflecting signature contents: "OFCLASS(?'a::c, c_class)" *) | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 1090 | fun class_triv thy c = | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1091 |   let val Cterm {thy_ref, t, maxidx, sorts, ...} =
 | 
| 19525 | 1092 |     cterm_of thy (Logic.mk_inclass (TVar (("'a", 0), [c]), Sign.certify_class thy c))
 | 
| 6368 | 1093 |       handle TERM (msg, _) => raise THM ("class_triv: " ^ msg, 0, []);
 | 
| 399 
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
 wenzelm parents: 
387diff
changeset | 1094 | in | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1095 |     Thm {thy_ref = thy_ref,
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1096 |       der = Pt.infer_derivs' I (false, Pt.PAxm ("ProtoPure.class_triv:" ^ c, t, SOME [])),
 | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 1097 | tags = [], | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1098 | maxidx = maxidx, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1099 | shyps = sorts, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1100 | hyps = [], | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1101 | tpairs = [], | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1102 | prop = t} | 
| 399 
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
 wenzelm parents: 
387diff
changeset | 1103 | end; | 
| 
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
 wenzelm parents: 
387diff
changeset | 1104 | |
| 19505 | 1105 | (*Internalize sort constraints of type variable*) | 
| 1106 | fun unconstrainT | |
| 1107 |     (Ctyp {thy_ref = thy_ref1, T, ...})
 | |
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 1108 |     (th as Thm {thy_ref = thy_ref2, der, maxidx, shyps, hyps, tpairs, prop, ...}) =
 | 
| 19505 | 1109 | let | 
| 1110 | val ((x, i), S) = Term.dest_TVar T handle TYPE _ => | |
| 1111 |       raise THM ("unconstrainT: not a type variable", 0, [th]);
 | |
| 1112 | val T' = TVar ((x, i), []); | |
| 20548 
8ef25fe585a8
renamed Term.map_term_types to Term.map_types (cf. Term.fold_types);
 wenzelm parents: 
20545diff
changeset | 1113 | val unconstrain = Term.map_types (Term.map_atyps (fn U => if U = T then T' else U)); | 
| 19505 | 1114 | val constraints = map (curry Logic.mk_inclass T') S; | 
| 1115 | in | |
| 1116 |     Thm {thy_ref = Theory.merge_refs (thy_ref1, thy_ref2),
 | |
| 1117 |       der = Pt.infer_derivs' I (false, Pt.PAxm ("ProtoPure.unconstrainT", prop, SOME [])),
 | |
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 1118 | tags = [], | 
| 19505 | 1119 | maxidx = Int.max (maxidx, i), | 
| 1120 | shyps = Sorts.remove_sort S shyps, | |
| 1121 | hyps = hyps, | |
| 1122 | tpairs = map (pairself unconstrain) tpairs, | |
| 1123 | prop = Logic.list_implies (constraints, unconstrain prop)} | |
| 1124 | end; | |
| 399 
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
 wenzelm parents: 
387diff
changeset | 1125 | |
| 6786 | 1126 | (* Replace all TFrees not fixed or in the hyps by new TVars *) | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 1127 | fun varifyT' fixed (Thm {thy_ref, der, maxidx, shyps, hyps, tpairs, prop, ...}) =
 | 
| 12500 | 1128 | let | 
| 15797 | 1129 | val tfrees = foldr add_term_tfrees fixed hyps; | 
| 13658 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 berghofe parents: 
13642diff
changeset | 1130 | val prop1 = attach_tpairs tpairs prop; | 
| 21116 | 1131 | val (al, prop2) = Type.varify tfrees prop1; | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1132 | val (ts, prop3) = Logic.strip_prems (length tpairs, [], prop2); | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1133 | in | 
| 18127 | 1134 |     (al, Thm {thy_ref = thy_ref,
 | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1135 | der = Pt.infer_derivs' (Pt.varify_proof prop tfrees) der, | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 1136 | tags = [], | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1137 | maxidx = Int.max (0, maxidx), | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1138 | shyps = shyps, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1139 | hyps = hyps, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1140 | tpairs = rev (map Logic.dest_equals ts), | 
| 18127 | 1141 | prop = prop3}) | 
| 0 | 1142 | end; | 
| 1143 | ||
| 18127 | 1144 | val varifyT = #2 o varifyT' []; | 
| 6786 | 1145 | |
| 0 | 1146 | (* Replace all TVars by new TFrees *) | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 1147 | fun freezeT (Thm {thy_ref, der, maxidx, shyps, hyps, tpairs, prop, ...}) =
 | 
| 13658 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 berghofe parents: 
13642diff
changeset | 1148 | let | 
| 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 berghofe parents: 
13642diff
changeset | 1149 | val prop1 = attach_tpairs tpairs prop; | 
| 16287 | 1150 | val prop2 = Type.freeze prop1; | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1151 | val (ts, prop3) = Logic.strip_prems (length tpairs, [], prop2); | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1152 | in | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1153 |     Thm {thy_ref = thy_ref,
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1154 | der = Pt.infer_derivs' (Pt.freezeT prop1) der, | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 1155 | tags = [], | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1156 | maxidx = maxidx_of_term prop2, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1157 | shyps = shyps, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1158 | hyps = hyps, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1159 | tpairs = rev (map Logic.dest_equals ts), | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1160 | prop = prop3} | 
| 1220 | 1161 | end; | 
| 0 | 1162 | |
| 1163 | ||
| 1164 | (*** Inference rules for tactics ***) | |
| 1165 | ||
| 1166 | (*Destruct proof state into constraints, other goals, goal(i), rest *) | |
| 13658 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 berghofe parents: 
13642diff
changeset | 1167 | fun dest_state (state as Thm{prop,tpairs,...}, i) =
 | 
| 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 berghofe parents: 
13642diff
changeset | 1168 | (case Logic.strip_prems(i, [], prop) of | 
| 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 berghofe parents: 
13642diff
changeset | 1169 | (B::rBs, C) => (tpairs, rev rBs, B, C) | 
| 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 berghofe parents: 
13642diff
changeset | 1170 |     | _ => raise THM("dest_state", i, [state]))
 | 
| 0 | 1171 |   handle TERM _ => raise THM("dest_state", i, [state]);
 | 
| 1172 | ||
| 309 | 1173 | (*Increment variables and parameters of orule as required for | 
| 18035 | 1174 | resolution with a goal.*) | 
| 1175 | fun lift_rule goal orule = | |
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1176 | let | 
| 18035 | 1177 |     val Cterm {t = gprop, T, maxidx = gmax, sorts, ...} = goal;
 | 
| 1178 | val inc = gmax + 1; | |
| 1179 | val lift_abs = Logic.lift_abs inc gprop; | |
| 1180 | val lift_all = Logic.lift_all inc gprop; | |
| 1181 |     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 | 1182 | val (As, B) = Logic.strip_horn prop; | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1183 | in | 
| 18035 | 1184 |     if T <> propT then raise THM ("lift_rule: the term must have type prop", 0, [])
 | 
| 1185 | else | |
| 1186 |       Thm {thy_ref = merge_thys1 goal orule,
 | |
| 1187 | der = Pt.infer_derivs' (Pt.lift_proof gprop inc prop) der, | |
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 1188 | tags = [], | 
| 18035 | 1189 | maxidx = maxidx + inc, | 
| 1190 | shyps = Sorts.union shyps sorts, (*sic!*) | |
| 1191 | hyps = hyps, | |
| 1192 | tpairs = map (pairself lift_abs) tpairs, | |
| 1193 | prop = Logic.list_implies (map lift_all As, lift_all B)} | |
| 0 | 1194 | end; | 
| 1195 | ||
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 1196 | fun incr_indexes i (thm as Thm {thy_ref, der, maxidx, shyps, hyps, tpairs, prop, ...}) =
 | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1197 |   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 | 1198 | else if i = 0 then thm | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1199 | else | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 1200 |     Thm {thy_ref = thy_ref,
 | 
| 16884 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 wenzelm parents: 
16847diff
changeset | 1201 | der = Pt.infer_derivs' | 
| 
1678a796b6b2
tuned instantiate (avoid subst_atomic, subst_atomic_types);
 wenzelm parents: 
16847diff
changeset | 1202 | (Pt.map_proof_terms (Logic.incr_indexes ([], i)) (Logic.incr_tvar i)) der, | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 1203 | tags = [], | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1204 | maxidx = maxidx + i, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1205 | shyps = shyps, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1206 | hyps = hyps, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1207 | tpairs = map (pairself (Logic.incr_indexes ([], i))) tpairs, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1208 | prop = Logic.incr_indexes ([], i) prop}; | 
| 10416 
5b33e732e459
- Moved rewriting functions to meta_simplifier.ML
 berghofe parents: 
10348diff
changeset | 1209 | |
| 0 | 1210 | (*Solve subgoal Bi of proof state B1...Bn/C by assumption. *) | 
| 1211 | fun assumption i state = | |
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1212 | let | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1213 |     val Thm {thy_ref, der, maxidx, shyps, hyps, prop, ...} = state;
 | 
| 16656 | 1214 | val thy = Theory.deref thy_ref; | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1215 | val (tpairs, Bs, Bi, C) = dest_state (state, i); | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1216 |     fun newth n (env as Envir.Envir {maxidx, ...}, tpairs) =
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1217 |       Thm {thy_ref = thy_ref,
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1218 | der = Pt.infer_derivs' | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1219 | ((if Envir.is_empty env then I else (Pt.norm_proof' env)) o | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1220 | Pt.assumption_proof Bs Bi n) der, | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 1221 | tags = [], | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1222 | maxidx = maxidx, | 
| 16656 | 1223 | shyps = may_insert_env_sorts thy env shyps, | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1224 | hyps = hyps, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1225 | tpairs = | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1226 | if Envir.is_empty env then tpairs | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1227 | else map (pairself (Envir.norm_term env)) tpairs, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1228 | prop = | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1229 | if Envir.is_empty env then (*avoid wasted normalizations*) | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1230 | Logic.list_implies (Bs, C) | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1231 | else (*normalize the new rule fully*) | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1232 | Envir.norm_term env (Logic.list_implies (Bs, C))}; | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1233 | fun addprfs [] _ = Seq.empty | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1234 | | addprfs ((t, u) :: apairs) n = Seq.make (fn () => Seq.pull | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1235 | (Seq.mapp (newth n) | 
| 16656 | 1236 | (Unify.unifiers (thy, Envir.empty maxidx, (t, u) :: tpairs)) | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1237 | (addprfs apairs (n + 1)))) | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1238 | in addprfs (Logic.assum_pairs (~1, Bi)) 1 end; | 
| 0 | 1239 | |
| 250 | 1240 | (*Solve subgoal Bi of proof state B1...Bn/C by assumption. | 
| 0 | 1241 | Checks if Bi's conclusion is alpha-convertible to one of its assumptions*) | 
| 1242 | fun eq_assumption i state = | |
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1243 | let | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1244 |     val Thm {thy_ref, der, maxidx, shyps, hyps, prop, ...} = state;
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1245 | val (tpairs, Bs, Bi, C) = dest_state (state, i); | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1246 | in | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1247 | (case find_index (op aconv) (Logic.assum_pairs (~1, Bi)) of | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1248 |       ~1 => raise THM ("eq_assumption", 0, [state])
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1249 | | n => | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1250 |         Thm {thy_ref = thy_ref,
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1251 | der = Pt.infer_derivs' (Pt.assumption_proof Bs Bi (n + 1)) der, | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 1252 | tags = [], | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1253 | maxidx = maxidx, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1254 | shyps = shyps, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1255 | hyps = hyps, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1256 | tpairs = tpairs, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1257 | prop = Logic.list_implies (Bs, C)}) | 
| 0 | 1258 | end; | 
| 1259 | ||
| 1260 | ||
| 2671 | 1261 | (*For rotate_tac: fast rotation of assumptions of subgoal i*) | 
| 1262 | fun rotate_rule k i state = | |
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1263 | let | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1264 |     val Thm {thy_ref, der, maxidx, shyps, hyps, prop, ...} = state;
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1265 | val (tpairs, Bs, Bi, C) = dest_state (state, i); | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1266 | val params = Term.strip_all_vars Bi | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1267 | and rest = Term.strip_all_body Bi; | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1268 | val asms = Logic.strip_imp_prems rest | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1269 | and concl = Logic.strip_imp_concl rest; | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1270 | val n = length asms; | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1271 | 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 | 1272 | val Bi' = | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1273 | if 0 = m orelse m = n then Bi | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1274 | else if 0 < m andalso m < n then | 
| 19012 | 1275 | let val (ps, qs) = chop m asms | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1276 | in list_all (params, Logic.list_implies (qs @ ps, concl)) end | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1277 |       else raise THM ("rotate_rule", k, [state]);
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1278 | in | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1279 |     Thm {thy_ref = thy_ref,
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1280 | der = Pt.infer_derivs' (Pt.rotate_proof Bs Bi m) der, | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 1281 | tags = [], | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1282 | maxidx = maxidx, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1283 | shyps = shyps, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1284 | hyps = hyps, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1285 | tpairs = tpairs, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1286 | prop = Logic.list_implies (Bs @ [Bi'], C)} | 
| 2671 | 1287 | end; | 
| 1288 | ||
| 1289 | ||
| 7248 
322151fe6f02
new primitive rule permute_prems to underlie defer_tac and rotate_prems
 paulson parents: 
7070diff
changeset | 1290 | (*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 | 1291 | unchanged. Does nothing if k=0 or if k equals n-j, where n is the | 
| 16656 | 1292 | number of premises. Useful with etac and underlies defer_tac*) | 
| 7248 
322151fe6f02
new primitive rule permute_prems to underlie defer_tac and rotate_prems
 paulson parents: 
7070diff
changeset | 1293 | fun permute_prems j k rl = | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1294 | let | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 1295 |     val Thm {thy_ref, der, maxidx, shyps, hyps, tpairs, prop, ...} = rl;
 | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1296 | val prems = Logic.strip_imp_prems prop | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1297 | and concl = Logic.strip_imp_concl prop; | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1298 | val moved_prems = List.drop (prems, j) | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1299 | and fixed_prems = List.take (prems, j) | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1300 |       handle Subscript => raise THM ("permute_prems: j", j, [rl]);
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1301 | val n_j = length moved_prems; | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1302 | 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 | 1303 | val prop' = | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1304 | if 0 = m orelse m = n_j then prop | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1305 | else if 0 < m andalso m < n_j then | 
| 19012 | 1306 | let val (ps, qs) = chop m moved_prems | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1307 | in Logic.list_implies (fixed_prems @ qs @ ps, concl) end | 
| 16725 | 1308 |       else raise THM ("permute_prems: k", k, [rl]);
 | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1309 | in | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1310 |     Thm {thy_ref = thy_ref,
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1311 | der = Pt.infer_derivs' (Pt.permute_prems_prf prems j m) der, | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 1312 | tags = [], | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1313 | maxidx = maxidx, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1314 | shyps = shyps, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1315 | hyps = hyps, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1316 | tpairs = tpairs, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1317 | prop = prop'} | 
| 7248 
322151fe6f02
new primitive rule permute_prems to underlie defer_tac and rotate_prems
 paulson parents: 
7070diff
changeset | 1318 | end; | 
| 
322151fe6f02
new primitive rule permute_prems to underlie defer_tac and rotate_prems
 paulson parents: 
7070diff
changeset | 1319 | |
| 
322151fe6f02
new primitive rule permute_prems to underlie defer_tac and rotate_prems
 paulson parents: 
7070diff
changeset | 1320 | |
| 0 | 1321 | (** User renaming of parameters in a subgoal **) | 
| 1322 | ||
| 1323 | (*Calls error rather than raising an exception because it is intended | |
| 1324 | for top-level use -- exception handling would not make sense here. | |
| 1325 | The names in cs, if distinct, are used for the innermost parameters; | |
| 17868 | 1326 | preceding parameters may be renamed to make all params distinct.*) | 
| 0 | 1327 | fun rename_params_rule (cs, i) state = | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1328 | let | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 1329 |     val Thm {thy_ref, der, tags, maxidx, shyps, hyps, ...} = state;
 | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1330 | val (tpairs, Bs, Bi, C) = dest_state (state, i); | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1331 | val iparams = map #1 (Logic.strip_params Bi); | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1332 | val short = length iparams - length cs; | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1333 | val newnames = | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1334 | if short < 0 then error "More names than abstractions!" | 
| 20071 
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
 wenzelm parents: 
20057diff
changeset | 1335 | else Name.variant_list cs (Library.take (short, iparams)) @ cs; | 
| 20330 | 1336 | val freenames = Term.fold_aterms (fn Free (x, _) => insert (op =) x | _ => I) Bi []; | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1337 | val newBi = Logic.list_rename_params (newnames, Bi); | 
| 250 | 1338 | in | 
| 21182 | 1339 | (case duplicates (op =) cs of | 
| 1340 |       a :: _ => (warning ("Can't rename.  Bound variables not distinct: " ^ a); state)
 | |
| 1341 | | [] => | |
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1342 | (case cs inter_string freenames of | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1343 |         a :: _ => (warning ("Can't rename.  Bound/Free variable clash: " ^ a); state)
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1344 | | [] => | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1345 |         Thm {thy_ref = thy_ref,
 | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1346 | der = der, | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 1347 | tags = tags, | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1348 | maxidx = maxidx, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1349 | shyps = shyps, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1350 | hyps = hyps, | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1351 | tpairs = tpairs, | 
| 21182 | 1352 | prop = Logic.list_implies (Bs @ [newBi], C)})) | 
| 0 | 1353 | end; | 
| 1354 | ||
| 12982 | 1355 | |
| 0 | 1356 | (*** Preservation of bound variable names ***) | 
| 1357 | ||
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 1358 | fun rename_boundvars pat obj (thm as Thm {thy_ref, der, tags, maxidx, shyps, hyps, tpairs, prop}) =
 | 
| 12982 | 1359 | (case Term.rename_abs pat obj prop of | 
| 15531 | 1360 | NONE => thm | 
| 1361 | | SOME prop' => Thm | |
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 1362 |       {thy_ref = thy_ref,
 | 
| 12982 | 1363 | der = der, | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 1364 | tags = tags, | 
| 12982 | 1365 | maxidx = maxidx, | 
| 1366 | hyps = hyps, | |
| 1367 | shyps = shyps, | |
| 13658 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 berghofe parents: 
13642diff
changeset | 1368 | tpairs = tpairs, | 
| 12982 | 1369 | prop = prop'}); | 
| 10416 
5b33e732e459
- Moved rewriting functions to meta_simplifier.ML
 berghofe parents: 
10348diff
changeset | 1370 | |
| 0 | 1371 | |
| 16656 | 1372 | (* strip_apply f (A, B) strips off all assumptions/parameters from A | 
| 0 | 1373 | introduced by lifting over B, and applies f to remaining part of A*) | 
| 1374 | fun strip_apply f = | |
| 1375 |   let fun strip(Const("==>",_)$ A1 $ B1,
 | |
| 250 | 1376 |                 Const("==>",_)$ _  $ B2) = implies $ A1 $ strip(B1,B2)
 | 
| 1377 |         | strip((c as Const("all",_)) $ Abs(a,T,t1),
 | |
| 1378 |                       Const("all",_)  $ Abs(_,_,t2)) = c$Abs(a,T,strip(t1,t2))
 | |
| 1379 | | strip(A,_) = f A | |
| 0 | 1380 | in strip end; | 
| 1381 | ||
| 1382 | (*Use the alist to rename all bound variables and some unknowns in a term | |
| 1383 | dpairs = current disagreement pairs; tpairs = permanent ones (flexflex); | |
| 1384 | Preserves unknowns in tpairs and on lhs of dpairs. *) | |
| 1385 | fun rename_bvs([],_,_,_) = I | |
| 1386 | | rename_bvs(al,dpairs,tpairs,B) = | |
| 20330 | 1387 | let | 
| 1388 | val add_var = fold_aterms (fn Var ((x, _), _) => insert (op =) x | _ => I); | |
| 1389 | val vids = [] | |
| 1390 | |> fold (add_var o fst) dpairs | |
| 1391 | |> fold (add_var o fst) tpairs | |
| 1392 | |> fold (add_var o snd) tpairs; | |
| 250 | 1393 | (*unknowns appearing elsewhere be preserved!*) | 
| 1394 | fun rename(t as Var((x,i),T)) = | |
| 20330 | 1395 | (case AList.lookup (op =) al x of | 
| 1396 | SOME y => | |
| 1397 | if member (op =) vids x orelse member (op =) vids y then t | |
| 1398 | else Var((y,i),T) | |
| 1399 | | NONE=> t) | |
| 0 | 1400 | | rename(Abs(x,T,t)) = | 
| 18944 | 1401 | Abs (the_default x (AList.lookup (op =) al x), T, rename t) | 
| 0 | 1402 | | rename(f$t) = rename f $ rename t | 
| 1403 | | rename(t) = t; | |
| 250 | 1404 | fun strip_ren Ai = strip_apply rename (Ai,B) | 
| 20330 | 1405 | in strip_ren end; | 
| 0 | 1406 | |
| 1407 | (*Function to rename bounds/unknowns in the argument, lifted over B*) | |
| 1408 | fun rename_bvars(dpairs, tpairs, B) = | |
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 1409 | rename_bvs(foldr Term.match_bvars [] dpairs, dpairs, tpairs, B); | 
| 0 | 1410 | |
| 1411 | ||
| 1412 | (*** RESOLUTION ***) | |
| 1413 | ||
| 721 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 lcp parents: 
678diff
changeset | 1414 | (** Lifting optimizations **) | 
| 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 lcp parents: 
678diff
changeset | 1415 | |
| 0 | 1416 | (*strip off pairs of assumptions/parameters in parallel -- they are | 
| 1417 | identical because of lifting*) | |
| 250 | 1418 | fun strip_assums2 (Const("==>", _) $ _ $ B1,
 | 
| 1419 |                    Const("==>", _) $ _ $ B2) = strip_assums2 (B1,B2)
 | |
| 0 | 1420 |   | strip_assums2 (Const("all",_)$Abs(a,T,t1),
 | 
| 250 | 1421 |                    Const("all",_)$Abs(_,_,t2)) =
 | 
| 0 | 1422 | let val (B1,B2) = strip_assums2 (t1,t2) | 
| 1423 | in (Abs(a,T,B1), Abs(a,T,B2)) end | |
| 1424 | | strip_assums2 BB = BB; | |
| 1425 | ||
| 1426 | ||
| 721 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 lcp parents: 
678diff
changeset | 1427 | (*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 | 1428 | fun norm_term_skip env 0 t = Envir.norm_term env t | 
| 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 lcp parents: 
678diff
changeset | 1429 |   | norm_term_skip env n (Const("all",_)$Abs(a,T,t)) =
 | 
| 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 lcp parents: 
678diff
changeset | 1430 |         let val Envir.Envir{iTs, ...} = env
 | 
| 15797 | 1431 | val T' = Envir.typ_subst_TVars iTs T | 
| 1238 | 1432 | (*Must instantiate types of parameters because they are flattened; | 
| 721 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 lcp parents: 
678diff
changeset | 1433 | this could be a NEW parameter*) | 
| 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 lcp parents: 
678diff
changeset | 1434 | in all T' $ Abs(a, T', norm_term_skip env n t) end | 
| 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 lcp parents: 
678diff
changeset | 1435 |   | norm_term_skip env n (Const("==>", _) $ A $ B) =
 | 
| 1238 | 1436 | implies $ A $ norm_term_skip env (n-1) B | 
| 721 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 lcp parents: 
678diff
changeset | 1437 | | norm_term_skip env n t = error"norm_term_skip: too few assumptions??"; | 
| 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 lcp parents: 
678diff
changeset | 1438 | |
| 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 lcp parents: 
678diff
changeset | 1439 | |
| 0 | 1440 | (*Composition of object rule r=(A1...Am/B) with proof state s=(B1...Bn/C) | 
| 250 | 1441 | Unifies B with Bi, replacing subgoal i (1 <= i <= n) | 
| 0 | 1442 | If match then forbid instantiations in proof state | 
| 1443 | If lifted then shorten the dpair using strip_assums2. | |
| 1444 | If eres_flg then simultaneously proves A1 by assumption. | |
| 250 | 1445 | nsubgoal is the number of new subgoals (written m above). | 
| 0 | 1446 | Curried so that resolution calls dest_state only once. | 
| 1447 | *) | |
| 4270 | 1448 | local exception COMPOSE | 
| 0 | 1449 | in | 
| 18486 | 1450 | fun bicompose_aux flatten match (state, (stpairs, Bs, Bi, C), lifted) | 
| 0 | 1451 | (eres_flg, orule, nsubgoal) = | 
| 1529 | 1452 |  let val Thm{der=sder, maxidx=smax, shyps=sshyps, hyps=shyps, ...} = state
 | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 1453 |      and Thm{der=rder, maxidx=rmax, shyps=rshyps, hyps=rhyps,
 | 
| 13658 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 berghofe parents: 
13642diff
changeset | 1454 | tpairs=rtpairs, prop=rprop,...} = orule | 
| 1529 | 1455 | (*How many hyps to skip over during normalization*) | 
| 21576 | 1456 | and nlift = Logic.count_prems (strip_all_body Bi) + (if eres_flg then ~1 else 0) | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1457 | val thy_ref = merge_thys2 state orule; | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 1458 | val thy = Theory.deref thy_ref; | 
| 0 | 1459 | (** Add new theorem with prop = '[| Bs; As |] ==> C' to thq **) | 
| 11518 | 1460 |      fun addth A (As, oldAs, rder', n) ((env as Envir.Envir {maxidx, ...}, tpairs), thq) =
 | 
| 250 | 1461 | let val normt = Envir.norm_term env; | 
| 1462 | (*perform minimal copying here by examining env*) | |
| 13658 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 berghofe parents: 
13642diff
changeset | 1463 | val (ntpairs, normp) = | 
| 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 berghofe parents: 
13642diff
changeset | 1464 | if Envir.is_empty env then (tpairs, (Bs @ As, C)) | 
| 250 | 1465 | else | 
| 1466 | let val ntps = map (pairself normt) tpairs | |
| 19861 | 1467 | in if Envir.above env smax then | 
| 1238 | 1468 | (*no assignments in state; normalize the rule only*) | 
| 1469 | if lifted | |
| 13658 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 berghofe parents: 
13642diff
changeset | 1470 | 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 | 1471 | else (ntps, (Bs @ map normt As, C)) | 
| 1529 | 1472 | else if match then raise COMPOSE | 
| 250 | 1473 | else (*normalize the new rule fully*) | 
| 13658 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 berghofe parents: 
13642diff
changeset | 1474 | (ntps, (map normt (Bs @ As), normt C)) | 
| 250 | 1475 | end | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1476 | val th = | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 1477 |              Thm{thy_ref = thy_ref,
 | 
| 11518 | 1478 | der = Pt.infer_derivs | 
| 1479 | ((if Envir.is_empty env then I | |
| 19861 | 1480 | else if Envir.above env smax then | 
| 11518 | 1481 | (fn f => fn der => f (Pt.norm_proof' env der)) | 
| 1482 | else | |
| 1483 | curry op oo (Pt.norm_proof' env)) | |
| 18486 | 1484 | (Pt.bicompose_proof flatten Bs oldAs As A n)) rder' sder, | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 1485 | tags = [], | 
| 2386 | 1486 | maxidx = maxidx, | 
| 16656 | 1487 | shyps = may_insert_env_sorts thy env (Sorts.union rshyps sshyps), | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1488 | hyps = union_hyps rhyps shyps, | 
| 13658 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 berghofe parents: 
13642diff
changeset | 1489 | tpairs = ntpairs, | 
| 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 berghofe parents: 
13642diff
changeset | 1490 | prop = Logic.list_implies normp} | 
| 19475 | 1491 | 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 | 1492 | val (rAs,B) = Logic.strip_prems(nsubgoal, [], rprop) | 
| 0 | 1493 |        handle TERM _ => raise THM("bicompose: rule", 0, [orule,state]);
 | 
| 1494 | (*Modify assumptions, deleting n-th if n>0 for e-resolution*) | |
| 1495 | fun newAs(As0, n, dpairs, tpairs) = | |
| 11518 | 1496 | let val (As1, rder') = | 
| 1497 | if !Logic.auto_rename orelse not lifted then (As0, rder) | |
| 1498 | else (map (rename_bvars(dpairs,tpairs,B)) As0, | |
| 1499 | Pt.infer_derivs' (Pt.map_proof_terms | |
| 1500 | (rename_bvars (dpairs, tpairs, Bound 0)) I) rder); | |
| 18486 | 1501 | in (map (if flatten then (Logic.flatten_params n) else I) As1, As1, rder', n) | 
| 250 | 1502 | handle TERM _ => | 
| 1503 |           raise THM("bicompose: 1st premise", 0, [orule])
 | |
| 0 | 1504 | end; | 
| 2147 | 1505 | val env = Envir.empty(Int.max(rmax,smax)); | 
| 0 | 1506 | val BBi = if lifted then strip_assums2(B,Bi) else (B,Bi); | 
| 1507 | val dpairs = BBi :: (rtpairs@stpairs); | |
| 1508 | (*elim-resolution: try each assumption in turn. Initially n=1*) | |
| 11518 | 1509 | fun tryasms (_, _, _, []) = Seq.empty | 
| 1510 | | tryasms (A, As, n, (t,u)::apairs) = | |
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 1511 | (case Seq.pull(Unify.unifiers(thy, env, (t,u)::dpairs)) of | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 1512 | NONE => tryasms (A, As, n+1, apairs) | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 1513 | | cell as SOME((_,tpairs),_) => | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 1514 | Seq.it_right (addth A (newAs(As, n, [BBi,(u,t)], tpairs))) | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 1515 | (Seq.make(fn()=> cell), | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 1516 | Seq.make(fn()=> Seq.pull (tryasms(A, As, n+1, apairs))))) | 
| 0 | 1517 |      fun eres [] = raise THM("bicompose: no premises", 0, [orule,state])
 | 
| 15531 | 1518 | | eres (A1::As) = tryasms(SOME A1, As, 1, Logic.assum_pairs(nlift+1,A1)) | 
| 0 | 1519 | (*ordinary resolution*) | 
| 15531 | 1520 | fun res(NONE) = Seq.empty | 
| 1521 | | res(cell as SOME((_,tpairs),_)) = | |
| 1522 | Seq.it_right (addth NONE (newAs(rev rAs, 0, [BBi], tpairs))) | |
| 4270 | 1523 | (Seq.make (fn()=> cell), Seq.empty) | 
| 0 | 1524 | in if eres_flg then eres(rev rAs) | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 1525 | else res(Seq.pull(Unify.unifiers(thy, env, dpairs))) | 
| 0 | 1526 | end; | 
| 7528 | 1527 | end; | 
| 0 | 1528 | |
| 18501 
915105af2e80
turned bicompose_no_flatten into compose_no_flatten, without elimination;
 wenzelm parents: 
18486diff
changeset | 1529 | fun compose_no_flatten match (orule, nsubgoal) i state = | 
| 
915105af2e80
turned bicompose_no_flatten into compose_no_flatten, without elimination;
 wenzelm parents: 
18486diff
changeset | 1530 | bicompose_aux false match (state, dest_state (state, i), false) (false, orule, nsubgoal); | 
| 0 | 1531 | |
| 18501 
915105af2e80
turned bicompose_no_flatten into compose_no_flatten, without elimination;
 wenzelm parents: 
18486diff
changeset | 1532 | fun bicompose match arg i state = | 
| 
915105af2e80
turned bicompose_no_flatten into compose_no_flatten, without elimination;
 wenzelm parents: 
18486diff
changeset | 1533 | bicompose_aux true match (state, dest_state (state,i), false) arg; | 
| 0 | 1534 | |
| 1535 | (*Quick test whether rule is resolvable with the subgoal with hyps Hs | |
| 1536 | and conclusion B. If eres_flg then checks 1st premise of rule also*) | |
| 1537 | fun could_bires (Hs, B, eres_flg, rule) = | |
| 16847 
8fc160b12e73
invoke_oracle: do not keep theory value, but theory_ref;
 wenzelm parents: 
16725diff
changeset | 1538 | let fun could_reshyp (A1::_) = exists (fn H => could_unify (A1, H)) Hs | 
| 250 | 1539 | | could_reshyp [] = false; (*no premise -- illegal*) | 
| 1540 | in could_unify(concl_of rule, B) andalso | |
| 1541 | (not eres_flg orelse could_reshyp (prems_of rule)) | |
| 0 | 1542 | end; | 
| 1543 | ||
| 1544 | (*Bi-resolution of a state with a list of (flag,rule) pairs. | |
| 1545 | Puts the rule above: rule/state. Renames vars in the rules. *) | |
| 250 | 1546 | fun biresolution match brules i state = | 
| 18035 | 1547 | let val (stpairs, Bs, Bi, C) = dest_state(state,i); | 
| 18145 | 1548 | val lift = lift_rule (cprem_of state i); | 
| 250 | 1549 | val B = Logic.strip_assums_concl Bi; | 
| 1550 | val Hs = Logic.strip_assums_hyp Bi; | |
| 22573 | 1551 | val compose = bicompose_aux true match (state, (stpairs, Bs, Bi, C), true); | 
| 4270 | 1552 | fun res [] = Seq.empty | 
| 250 | 1553 | | res ((eres_flg, rule)::brules) = | 
| 13642 
a3d97348ceb6
added failure trace information to pattern unification
 nipkow parents: 
13629diff
changeset | 1554 | if !Pattern.trace_unify_fail orelse | 
| 
a3d97348ceb6
added failure trace information to pattern unification
 nipkow parents: 
13629diff
changeset | 1555 | could_bires (Hs, B, eres_flg, rule) | 
| 4270 | 1556 | then Seq.make (*delay processing remainder till needed*) | 
| 22573 | 1557 | (fn()=> SOME(compose (eres_flg, lift rule, nprems_of rule), | 
| 250 | 1558 | res brules)) | 
| 1559 | else res brules | |
| 4270 | 1560 | in Seq.flat (res brules) end; | 
| 0 | 1561 | |
| 1562 | ||
| 2509 | 1563 | (*** Oracles ***) | 
| 1564 | ||
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 1565 | fun invoke_oracle_i thy1 name = | 
| 3812 | 1566 | let | 
| 1567 | val oracle = | |
| 22685 | 1568 | (case Symtab.lookup (Theory.oracle_table thy1) name of | 
| 15531 | 1569 |         NONE => raise THM ("Unknown oracle: " ^ name, 0, [])
 | 
| 1570 | | SOME (f, _) => f); | |
| 16847 
8fc160b12e73
invoke_oracle: do not keep theory value, but theory_ref;
 wenzelm parents: 
16725diff
changeset | 1571 | val thy_ref1 = Theory.self_ref thy1; | 
| 3812 | 1572 | in | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 1573 | fn (thy2, data) => | 
| 3812 | 1574 | let | 
| 16847 
8fc160b12e73
invoke_oracle: do not keep theory value, but theory_ref;
 wenzelm parents: 
16725diff
changeset | 1575 | val thy' = Theory.merge (Theory.deref thy_ref1, thy2); | 
| 18969 | 1576 | val (prop, T, maxidx) = Sign.certify_term thy' (oracle (thy', data)); | 
| 3812 | 1577 | in | 
| 1578 | if T <> propT then | |
| 1579 |           raise THM ("Oracle's result must have type prop: " ^ name, 0, [])
 | |
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1580 | else | 
| 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1581 |           Thm {thy_ref = Theory.self_ref thy',
 | 
| 11518 | 1582 | der = (true, Pt.oracle_proof name prop), | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21576diff
changeset | 1583 | tags = [], | 
| 3812 | 1584 | maxidx = maxidx, | 
| 16656 | 1585 | shyps = may_insert_term_sorts thy' prop [], | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 1586 | hyps = [], | 
| 13658 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 berghofe parents: 
13642diff
changeset | 1587 | tpairs = [], | 
| 16601 
ee8eefade568
more efficient treatment of shyps and hyps (use ordered lists);
 wenzelm parents: 
16425diff
changeset | 1588 | prop = prop} | 
| 3812 | 1589 | end | 
| 1590 | end; | |
| 1591 | ||
| 15672 | 1592 | fun invoke_oracle thy = | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16352diff
changeset | 1593 | invoke_oracle_i thy o NameSpace.intern (Theory.oracle_space thy); | 
| 15672 | 1594 | |
| 22237 
bb9b1c8a8a95
old-fashioned abstype ctyp/cterm/thm prevents accidental equality tests;
 wenzelm parents: 
21975diff
changeset | 1595 | |
| 
bb9b1c8a8a95
old-fashioned abstype ctyp/cterm/thm prevents accidental equality tests;
 wenzelm parents: 
21975diff
changeset | 1596 | end; | 
| 
bb9b1c8a8a95
old-fashioned abstype ctyp/cterm/thm prevents accidental equality tests;
 wenzelm parents: 
21975diff
changeset | 1597 | end; | 
| 
bb9b1c8a8a95
old-fashioned abstype ctyp/cterm/thm prevents accidental equality tests;
 wenzelm parents: 
21975diff
changeset | 1598 | end; | 
| 0 | 1599 | end; | 
| 1503 | 1600 | |
| 6089 | 1601 | structure BasicThm: BASIC_THM = Thm; | 
| 1602 | open BasicThm; |