| author | Fabian Huch <huch@in.tum.de> | 
| Fri, 28 Jun 2024 11:51:46 +0200 | |
| changeset 80420 | 80e10a1aa431 | 
| parent 78800 | 0b3700d31758 | 
| child 82965 | 8142462f0883 | 
| permissions | -rw-r--r-- | 
| 5982 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 1 | (* Title: Provers/Arith/fast_lin_arith.ML | 
| 31510 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 2 | Author: Tobias Nipkow and Tjark Weber and Sascha Boehme | 
| 6102 | 3 | |
| 46709 
65a9b30bff00
clarified prems_lin_arith_tac, with subtle change of semantics: structured prems are inserted as well;
 wenzelm parents: 
44654diff
changeset | 4 | A generic linear arithmetic package. | 
| 6102 | 5 | |
| 24076 | 6 | Only take premises and conclusions into account that are already | 
| 7 | (negated) (in)equations. lin_arith_simproc tries to prove or disprove | |
| 8 | the term. | |
| 5982 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 9 | *) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 10 | |
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 11 | (*** Data needed for setting up the linear arithmetic package ***) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 12 | |
| 6102 | 13 | signature LIN_ARITH_LOGIC = | 
| 14 | sig | |
| 20276 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 15 | val conjI : thm (* P ==> Q ==> P & Q *) | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 16 | val ccontr : thm (* (~ P ==> False) ==> P *) | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 17 | val notI : thm (* (P ==> False) ==> ~ P *) | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 18 | val not_lessD : thm (* ~(m < n) ==> n <= m *) | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 19 | val not_leD : thm (* ~(m <= n) ==> n < m *) | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 20 | val sym : thm (* x = y ==> y = x *) | 
| 31510 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 21 | val trueI : thm (* True *) | 
| 20276 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 22 | val mk_Eq : thm -> thm | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 23 | val atomize : thm -> thm list | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 24 | val mk_Trueprop : term -> term | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 25 | val neg_prop : term -> term | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 26 | val is_False : thm -> bool | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 27 | val is_nat : typ list * term -> bool | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 28 | val mk_nat_thm : theory -> term -> thm | 
| 6102 | 29 | end; | 
| 30 | (* | |
| 31 | mk_Eq(~in) = `in == False' | |
| 32 | mk_Eq(in) = `in == True' | |
| 33 | where `in' is an (in)equality. | |
| 34 | ||
| 23190 | 35 | neg_prop(t) = neg if t is wrapped up in Trueprop and neg is the | 
| 36 | (logically) negated version of t (again wrapped up in Trueprop), | |
| 37 | where the negation of a negative term is the term itself (no | |
| 38 |   double negation!); raises TERM ("neg_prop", [t]) if t is not of
 | |
| 39 | the form 'Trueprop $ _' | |
| 6128 | 40 | |
| 41 | is_nat(parameter-types,t) = t:nat | |
| 42 | mk_nat_thm(t) = "0 <= t" | |
| 6102 | 43 | *) | 
| 44 | ||
| 5982 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 45 | signature LIN_ARITH_DATA = | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 46 | sig | 
| 24076 | 47 | (*internal representation of linear (in-)equations:*) | 
| 26945 | 48 | type decomp = (term * Rat.rat) list * Rat.rat * string * (term * Rat.rat) list * Rat.rat * bool | 
| 49 | val decomp: Proof.context -> term -> decomp option | |
| 24076 | 50 | val domain_is_nat: term -> bool | 
| 51 | ||
| 66035 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 52 | (*abstraction for proof replay*) | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 53 | val abstract_arith: term -> (term * term) list * Proof.context -> | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 54 | term * ((term * term) list * Proof.context) | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 55 | val abstract: term -> (term * term) list * Proof.context -> | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 56 | term * ((term * term) list * Proof.context) | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 57 | |
| 24076 | 58 | (*preprocessing, performed on a representation of subgoals as list of premises:*) | 
| 59 | val pre_decomp: Proof.context -> typ list * term list -> (typ list * term list) list | |
| 60 | ||
| 61 | (*preprocessing, performed on the goal -- must do the same as 'pre_decomp':*) | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 62 | val pre_tac: Proof.context -> int -> tactic | 
| 24076 | 63 | |
| 64 | (*the limit on the number of ~= allowed; because each ~= is split | |
| 65 | into two cases, this can lead to an explosion*) | |
| 44654 | 66 | val neq_limit: int Config.T | 
| 43607 
119767e1ccb4
linarith counterexamples now provide only valuations for variables (which should restrict the number of linarith trace messages);
 boehmes parents: 
43568diff
changeset | 67 | |
| 44654 | 68 | val trace: bool Config.T | 
| 5982 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 69 | end; | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 70 | (* | 
| 7551 
8e934d1a9ac6
Now distinguishes discrete from non-distrete types.
 nipkow parents: 
6128diff
changeset | 71 | decomp(`x Rel y') should yield (p,i,Rel,q,j,d) | 
| 5982 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 72 | where Rel is one of "<", "~<", "<=", "~<=" and "=" and | 
| 20217 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19618diff
changeset | 73 | p (q, respectively) is the decomposition of the sum term x | 
| 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19618diff
changeset | 74 | (y, respectively) into a list of summand * multiplicity | 
| 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19618diff
changeset | 75 | pairs and a constant summand and d indicates if the domain | 
| 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19618diff
changeset | 76 | is discrete. | 
| 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19618diff
changeset | 77 | |
| 20276 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 78 | domain_is_nat(`x Rel y') t should yield true iff x is of type "nat". | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 79 | |
| 20217 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19618diff
changeset | 80 | The relationship between pre_decomp and pre_tac is somewhat tricky. The | 
| 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19618diff
changeset | 81 | internal representation of a subgoal and the corresponding theorem must | 
| 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19618diff
changeset | 82 | be modified by pre_decomp (pre_tac, resp.) in a corresponding way. See | 
| 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19618diff
changeset | 83 | the comment for split_items below. (This is even necessary for eta- and | 
| 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19618diff
changeset | 84 | beta-equivalent modifications, as some of the lin. arith. code is not | 
| 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19618diff
changeset | 85 | insensitive to them.) | 
| 5982 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 86 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 87 | Simpset must reduce contradictory <= to False. | 
| 5982 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 88 | It should also cancel common summands to keep <= reduced; | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 89 | otherwise <= can grow to massive proportions. | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 90 | *) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 91 | |
| 6062 | 92 | signature FAST_LIN_ARITH = | 
| 93 | sig | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 94 | val prems_lin_arith_tac: Proof.context -> int -> tactic | 
| 59656 | 95 | val lin_arith_tac: Proof.context -> int -> tactic | 
| 78800 | 96 | val lin_arith_simproc: Simplifier.proc | 
| 38763 
283f1f9969ba
Fast_Lin_Arith.number_of: more conventional merge that prefers the left side -- note that former ordering wrt. serial numbers makes it depend on accidental load order;
 wenzelm parents: 
38762diff
changeset | 97 | val map_data: | 
| 
283f1f9969ba
Fast_Lin_Arith.number_of: more conventional merge that prefers the left side -- note that former ordering wrt. serial numbers makes it depend on accidental load order;
 wenzelm parents: 
38762diff
changeset | 98 |     ({add_mono_thms: thm list, mult_mono_thms: thm list, inj_thms: thm list,
 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 99 | lessD: thm list, neqE: thm list, simpset: simpset, | 
| 59996 | 100 | number_of: (Proof.context -> typ -> int -> cterm) option} -> | 
| 38763 
283f1f9969ba
Fast_Lin_Arith.number_of: more conventional merge that prefers the left side -- note that former ordering wrt. serial numbers makes it depend on accidental load order;
 wenzelm parents: 
38762diff
changeset | 101 |      {add_mono_thms: thm list, mult_mono_thms: thm list, inj_thms: thm list,
 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 102 | lessD: thm list, neqE: thm list, simpset: simpset, | 
| 59996 | 103 | number_of: (Proof.context -> typ -> int -> cterm) option}) -> | 
| 38763 
283f1f9969ba
Fast_Lin_Arith.number_of: more conventional merge that prefers the left side -- note that former ordering wrt. serial numbers makes it depend on accidental load order;
 wenzelm parents: 
38762diff
changeset | 104 | Context.generic -> Context.generic | 
| 38762 
996afaa9254a
slightly more abstract data handling in Fast_Lin_Arith;
 wenzelm parents: 
38052diff
changeset | 105 | val add_inj_thms: thm list -> Context.generic -> Context.generic | 
| 
996afaa9254a
slightly more abstract data handling in Fast_Lin_Arith;
 wenzelm parents: 
38052diff
changeset | 106 | val add_lessD: thm -> Context.generic -> Context.generic | 
| 
996afaa9254a
slightly more abstract data handling in Fast_Lin_Arith;
 wenzelm parents: 
38052diff
changeset | 107 | val add_simps: thm list -> Context.generic -> Context.generic | 
| 
996afaa9254a
slightly more abstract data handling in Fast_Lin_Arith;
 wenzelm parents: 
38052diff
changeset | 108 | val add_simprocs: simproc list -> Context.generic -> Context.generic | 
| 59996 | 109 | val set_number_of: (Proof.context -> typ -> int -> cterm) -> Context.generic -> Context.generic | 
| 6062 | 110 | end; | 
| 111 | ||
| 24076 | 112 | functor Fast_Lin_Arith | 
| 113 | (structure LA_Logic: LIN_ARITH_LOGIC and LA_Data: LIN_ARITH_DATA): FAST_LIN_ARITH = | |
| 5982 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 114 | struct | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 115 | |
| 9420 | 116 | |
| 117 | (** theory data **) | |
| 118 | ||
| 33519 | 119 | structure Data = Generic_Data | 
| 22846 | 120 | ( | 
| 24076 | 121 | type T = | 
| 122 |    {add_mono_thms: thm list,
 | |
| 123 | mult_mono_thms: thm list, | |
| 124 | inj_thms: thm list, | |
| 125 | lessD: thm list, | |
| 126 | neqE: thm list, | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 127 | simpset: simpset, | 
| 59996 | 128 | number_of: (Proof.context -> typ -> int -> cterm) option}; | 
| 9420 | 129 | |
| 38763 
283f1f9969ba
Fast_Lin_Arith.number_of: more conventional merge that prefers the left side -- note that former ordering wrt. serial numbers makes it depend on accidental load order;
 wenzelm parents: 
38762diff
changeset | 130 | val empty : T = | 
| 
283f1f9969ba
Fast_Lin_Arith.number_of: more conventional merge that prefers the left side -- note that former ordering wrt. serial numbers makes it depend on accidental load order;
 wenzelm parents: 
38762diff
changeset | 131 |    {add_mono_thms = [], mult_mono_thms = [], inj_thms = [],
 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 132 | lessD = [], neqE = [], simpset = empty_ss, | 
| 38763 
283f1f9969ba
Fast_Lin_Arith.number_of: more conventional merge that prefers the left side -- note that former ordering wrt. serial numbers makes it depend on accidental load order;
 wenzelm parents: 
38762diff
changeset | 133 | number_of = NONE}; | 
| 33519 | 134 | fun merge | 
| 38763 
283f1f9969ba
Fast_Lin_Arith.number_of: more conventional merge that prefers the left side -- note that former ordering wrt. serial numbers makes it depend on accidental load order;
 wenzelm parents: 
38762diff
changeset | 135 |     ({add_mono_thms = add_mono_thms1, mult_mono_thms = mult_mono_thms1, inj_thms = inj_thms1,
 | 
| 
283f1f9969ba
Fast_Lin_Arith.number_of: more conventional merge that prefers the left side -- note that former ordering wrt. serial numbers makes it depend on accidental load order;
 wenzelm parents: 
38762diff
changeset | 136 | lessD = lessD1, neqE = neqE1, simpset = simpset1, number_of = number_of1}, | 
| 
283f1f9969ba
Fast_Lin_Arith.number_of: more conventional merge that prefers the left side -- note that former ordering wrt. serial numbers makes it depend on accidental load order;
 wenzelm parents: 
38762diff
changeset | 137 |      {add_mono_thms = add_mono_thms2, mult_mono_thms = mult_mono_thms2, inj_thms = inj_thms2,
 | 
| 
283f1f9969ba
Fast_Lin_Arith.number_of: more conventional merge that prefers the left side -- note that former ordering wrt. serial numbers makes it depend on accidental load order;
 wenzelm parents: 
38762diff
changeset | 138 | lessD = lessD2, neqE = neqE2, simpset = simpset2, number_of = number_of2}) : T = | 
| 24039 
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
23577diff
changeset | 139 |     {add_mono_thms = Thm.merge_thms (add_mono_thms1, add_mono_thms2),
 | 
| 
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
23577diff
changeset | 140 | mult_mono_thms = Thm.merge_thms (mult_mono_thms1, mult_mono_thms2), | 
| 
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
23577diff
changeset | 141 | inj_thms = Thm.merge_thms (inj_thms1, inj_thms2), | 
| 
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
23577diff
changeset | 142 | lessD = Thm.merge_thms (lessD1, lessD2), | 
| 
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
23577diff
changeset | 143 | neqE = Thm.merge_thms (neqE1, neqE2), | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 144 | simpset = merge_ss (simpset1, simpset2), | 
| 41493 | 145 | number_of = merge_options (number_of1, number_of2)}; | 
| 22846 | 146 | ); | 
| 9420 | 147 | |
| 148 | val map_data = Data.map; | |
| 24076 | 149 | val get_data = Data.get o Context.Proof; | 
| 9420 | 150 | |
| 67649 | 151 | fun get_neqE ctxt = map (Thm.transfer' ctxt) (#neqE (get_data ctxt)); | 
| 61097 | 152 | |
| 38762 
996afaa9254a
slightly more abstract data handling in Fast_Lin_Arith;
 wenzelm parents: 
38052diff
changeset | 153 | fun map_inj_thms f {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset, number_of} =
 | 
| 
996afaa9254a
slightly more abstract data handling in Fast_Lin_Arith;
 wenzelm parents: 
38052diff
changeset | 154 |   {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms, inj_thms = f inj_thms,
 | 
| 
996afaa9254a
slightly more abstract data handling in Fast_Lin_Arith;
 wenzelm parents: 
38052diff
changeset | 155 | lessD = lessD, neqE = neqE, simpset = simpset, number_of = number_of}; | 
| 
996afaa9254a
slightly more abstract data handling in Fast_Lin_Arith;
 wenzelm parents: 
38052diff
changeset | 156 | |
| 
996afaa9254a
slightly more abstract data handling in Fast_Lin_Arith;
 wenzelm parents: 
38052diff
changeset | 157 | fun map_lessD f {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset, number_of} =
 | 
| 
996afaa9254a
slightly more abstract data handling in Fast_Lin_Arith;
 wenzelm parents: 
38052diff
changeset | 158 |   {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms, inj_thms = inj_thms,
 | 
| 
996afaa9254a
slightly more abstract data handling in Fast_Lin_Arith;
 wenzelm parents: 
38052diff
changeset | 159 | lessD = f lessD, neqE = neqE, simpset = simpset, number_of = number_of}; | 
| 
996afaa9254a
slightly more abstract data handling in Fast_Lin_Arith;
 wenzelm parents: 
38052diff
changeset | 160 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 161 | fun map_simpset f context = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 162 |   map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset, number_of} =>
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 163 |     {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms, inj_thms = inj_thms,
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 164 | lessD = lessD, neqE = neqE, simpset = simpset_map (Context.proof_of context) f simpset, | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 165 | number_of = number_of}) context; | 
| 38762 
996afaa9254a
slightly more abstract data handling in Fast_Lin_Arith;
 wenzelm parents: 
38052diff
changeset | 166 | |
| 61097 | 167 | fun add_inj_thms thms = map_data (map_inj_thms (append (map Thm.trim_context thms))); | 
| 168 | fun add_lessD thm = map_data (map_lessD (fn thms => thms @ [Thm.trim_context thm])); | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 169 | fun add_simps thms = map_simpset (fn ctxt => ctxt addsimps thms); | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 170 | fun add_simprocs procs = map_simpset (fn ctxt => ctxt addsimprocs procs); | 
| 38762 
996afaa9254a
slightly more abstract data handling in Fast_Lin_Arith;
 wenzelm parents: 
38052diff
changeset | 171 | |
| 38763 
283f1f9969ba
Fast_Lin_Arith.number_of: more conventional merge that prefers the left side -- note that former ordering wrt. serial numbers makes it depend on accidental load order;
 wenzelm parents: 
38762diff
changeset | 172 | fun set_number_of f = | 
| 
283f1f9969ba
Fast_Lin_Arith.number_of: more conventional merge that prefers the left side -- note that former ordering wrt. serial numbers makes it depend on accidental load order;
 wenzelm parents: 
38762diff
changeset | 173 |   map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset, ...} =>
 | 
| 
283f1f9969ba
Fast_Lin_Arith.number_of: more conventional merge that prefers the left side -- note that former ordering wrt. serial numbers makes it depend on accidental load order;
 wenzelm parents: 
38762diff
changeset | 174 |    {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms, inj_thms = inj_thms,
 | 
| 
283f1f9969ba
Fast_Lin_Arith.number_of: more conventional merge that prefers the left side -- note that former ordering wrt. serial numbers makes it depend on accidental load order;
 wenzelm parents: 
38762diff
changeset | 175 | lessD = lessD, neqE = neqE, simpset = simpset, number_of = SOME f}); | 
| 
283f1f9969ba
Fast_Lin_Arith.number_of: more conventional merge that prefers the left side -- note that former ordering wrt. serial numbers makes it depend on accidental load order;
 wenzelm parents: 
38762diff
changeset | 176 | |
| 
283f1f9969ba
Fast_Lin_Arith.number_of: more conventional merge that prefers the left side -- note that former ordering wrt. serial numbers makes it depend on accidental load order;
 wenzelm parents: 
38762diff
changeset | 177 | fun number_of ctxt = | 
| 61097 | 178 | (case get_data ctxt of | 
| 59996 | 179 |     {number_of = SOME f, ...} => f ctxt
 | 
| 38763 
283f1f9969ba
Fast_Lin_Arith.number_of: more conventional merge that prefers the left side -- note that former ordering wrt. serial numbers makes it depend on accidental load order;
 wenzelm parents: 
38762diff
changeset | 180 |   | _ => fn _ => fn _ => raise CTERM ("number_of", []));
 | 
| 
283f1f9969ba
Fast_Lin_Arith.number_of: more conventional merge that prefers the left side -- note that former ordering wrt. serial numbers makes it depend on accidental load order;
 wenzelm parents: 
38762diff
changeset | 181 | |
| 9420 | 182 | |
| 183 | ||
| 5982 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 184 | (*** A fast decision procedure ***) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 185 | (*** Code ported from HOL Light ***) | 
| 6056 | 186 | (* possible optimizations: | 
| 187 | use (var,coeff) rep or vector rep tp save space; | |
| 188 | treat non-negative atoms separately rather than adding 0 <= atom | |
| 189 | *) | |
| 5982 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 190 | |
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 191 | datatype lineq_type = Eq | Le | Lt; | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 192 | |
| 6056 | 193 | datatype injust = Asm of int | 
| 194 | | Nat of int (* index of atom *) | |
| 6128 | 195 | | LessD of injust | 
| 196 | | NotLessD of injust | |
| 197 | | NotLeD of injust | |
| 7551 
8e934d1a9ac6
Now distinguishes discrete from non-distrete types.
 nipkow parents: 
6128diff
changeset | 198 | | NotLeDD of injust | 
| 24630 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
24112diff
changeset | 199 | | Multiplied of int * injust | 
| 5982 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 200 | | Added of injust * injust; | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 201 | |
| 24630 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
24112diff
changeset | 202 | datatype lineq = Lineq of int * lineq_type * int list * injust; | 
| 5982 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 203 | |
| 13498 | 204 | (* ------------------------------------------------------------------------- *) | 
| 205 | (* Finding a (counter) example from the trace of a failed elimination *) | |
| 206 | (* ------------------------------------------------------------------------- *) | |
| 207 | (* Examples are represented as rational numbers, *) | |
| 208 | (* Dont blame John Harrison for this code - it is entirely mine. TN *) | |
| 209 | ||
| 210 | exception NoEx; | |
| 211 | ||
| 14372 
51ddf8963c95
Finally fixed the counterexample finder. Can now deal with < on real.
 nipkow parents: 
14360diff
changeset | 212 | (* Coding: (i,true,cs) means i <= cs and (i,false,cs) means i < cs. | 
| 
51ddf8963c95
Finally fixed the counterexample finder. Can now deal with < on real.
 nipkow parents: 
14360diff
changeset | 213 | In general, true means the bound is included, false means it is excluded. | 
| 
51ddf8963c95
Finally fixed the counterexample finder. Can now deal with < on real.
 nipkow parents: 
14360diff
changeset | 214 | Need to know if it is a lower or upper bound for unambiguous interpretation! | 
| 
51ddf8963c95
Finally fixed the counterexample finder. Can now deal with < on real.
 nipkow parents: 
14360diff
changeset | 215 | *) | 
| 
51ddf8963c95
Finally fixed the counterexample finder. Can now deal with < on real.
 nipkow parents: 
14360diff
changeset | 216 | |
| 13498 | 217 | (* ------------------------------------------------------------------------- *) | 
| 23197 | 218 | (* End of counterexample finder. The actual decision procedure starts here. *) | 
| 13498 | 219 | (* ------------------------------------------------------------------------- *) | 
| 220 | ||
| 5982 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 221 | (* ------------------------------------------------------------------------- *) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 222 | (* Calculate new (in)equality type after addition. *) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 223 | (* ------------------------------------------------------------------------- *) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 224 | |
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 225 | fun find_add_type(Eq,x) = x | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 226 | | find_add_type(x,Eq) = x | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 227 | | find_add_type(_,Lt) = Lt | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 228 | | find_add_type(Lt,_) = Lt | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 229 | | find_add_type(Le,Le) = Le; | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 230 | |
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 231 | (* ------------------------------------------------------------------------- *) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 232 | (* Multiply out an (in)equation. *) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 233 | (* ------------------------------------------------------------------------- *) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 234 | |
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 235 | fun multiply_ineq n (i as Lineq(k,ty,l,just)) = | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 236 | if n = 1 then i | 
| 40316 
665862241968
replaced ancient sys_error by raise Fail, assuming that the latter is not handled specifically by the environment;
 wenzelm parents: 
38763diff
changeset | 237 | else if n = 0 andalso ty = Lt then raise Fail "multiply_ineq" | 
| 
665862241968
replaced ancient sys_error by raise Fail, assuming that the latter is not handled specifically by the environment;
 wenzelm parents: 
38763diff
changeset | 238 | else if n < 0 andalso (ty=Le orelse ty=Lt) then raise Fail "multiply_ineq" | 
| 33002 | 239 | else Lineq (n * k, ty, map (Integer.mult n) l, Multiplied (n, just)); | 
| 5982 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 240 | |
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 241 | (* ------------------------------------------------------------------------- *) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 242 | (* Add together (in)equations. *) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 243 | (* ------------------------------------------------------------------------- *) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 244 | |
| 38052 | 245 | fun add_ineq (Lineq (k1,ty1,l1,just1)) (Lineq (k2,ty2,l2,just2)) = | 
| 33002 | 246 | let val l = map2 Integer.add l1 l2 | 
| 5982 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 247 | in Lineq(k1+k2,find_add_type(ty1,ty2),l,Added(just1,just2)) end; | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 248 | |
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 249 | (* ------------------------------------------------------------------------- *) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 250 | (* Elimination of variable between a single pair of (in)equations. *) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 251 | (* If they're both inequalities, 1st coefficient must be +ve, 2nd -ve. *) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 252 | (* ------------------------------------------------------------------------- *) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 253 | |
| 49387 | 254 | fun elim_var v (i1 as Lineq(_,ty1,l1,_)) (i2 as Lineq(_,ty2,l2,_)) = | 
| 23063 | 255 | let val c1 = nth l1 v and c2 = nth l2 v | 
| 63227 
d3ed7f00e818
Integer.lcm normalizes the sign as in HOL/GCD.thy;
 wenzelm parents: 
63201diff
changeset | 256 | val m = Integer.lcm c1 c2 | 
| 5982 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 257 | val m1 = m div (abs c1) and m2 = m div (abs c2) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 258 | val (n1,n2) = | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 259 | if (c1 >= 0) = (c2 >= 0) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 260 | then if ty1 = Eq then (~m1,m2) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 261 | else if ty2 = Eq then (m1,~m2) | 
| 40316 
665862241968
replaced ancient sys_error by raise Fail, assuming that the latter is not handled specifically by the environment;
 wenzelm parents: 
38763diff
changeset | 262 | else raise Fail "elim_var" | 
| 5982 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 263 | else (m1,m2) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 264 | val (p1,p2) = if ty1=Eq andalso ty2=Eq andalso (n1 = ~1 orelse n2 = ~1) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 265 | then (~n1,~n2) else (n1,n2) | 
| 31510 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 266 | in add_ineq (multiply_ineq p1 i1) (multiply_ineq p2 i2) end; | 
| 5982 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 267 | |
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 268 | (* ------------------------------------------------------------------------- *) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 269 | (* The main refutation-finding code. *) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 270 | (* ------------------------------------------------------------------------- *) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 271 | |
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 272 | fun is_trivial (Lineq(_,_,l,_)) = forall (fn i => i=0) l; | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 273 | |
| 38052 | 274 | fun is_contradictory (Lineq(k,ty,_,_)) = | 
| 5982 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 275 | case ty of Eq => k <> 0 | Le => k > 0 | Lt => k >= 0; | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 276 | |
| 24630 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
24112diff
changeset | 277 | fun calc_blowup l = | 
| 33317 | 278 | let val (p,n) = List.partition (curry (op <) 0) (filter (curry (op <>) 0) l) | 
| 24630 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
24112diff
changeset | 279 | in length p * length n end; | 
| 5982 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 280 | |
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 281 | (* ------------------------------------------------------------------------- *) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 282 | (* Main elimination code: *) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 283 | (* *) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 284 | (* (1) Looks for immediate solutions (false assertions with no variables). *) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 285 | (* *) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 286 | (* (2) If there are any equations, picks a variable with the lowest absolute *) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 287 | (* coefficient in any of them, and uses it to eliminate. *) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 288 | (* *) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 289 | (* (3) Otherwise, chooses a variable in the inequality to minimize the *) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 290 | (* blowup (number of consequences generated) and eliminates it. *) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 291 | (* ------------------------------------------------------------------------- *) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 292 | |
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 293 | fun extract_first p = | 
| 31510 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 294 | let | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 295 | fun extract xs (y::ys) = if p y then (y, xs @ ys) else extract (y::xs) ys | 
| 66035 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 296 | | extract _ [] = raise List.Empty | 
| 5982 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 297 | in extract [] end; | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 298 | |
| 44654 | 299 | fun print_ineqs ctxt ineqs = | 
| 300 | if Config.get ctxt LA_Data.trace then | |
| 12262 | 301 |      tracing(cat_lines(""::map (fn Lineq(c,t,l,_) =>
 | 
| 24630 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
24112diff
changeset | 302 | string_of_int c ^ | 
| 9073 | 303 | (case t of Eq => " = " | Lt=> " < " | Le => " <= ") ^ | 
| 24630 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
24112diff
changeset | 304 | commas(map string_of_int l)) ineqs)) | 
| 9073 | 305 | else (); | 
| 6056 | 306 | |
| 13498 | 307 | type history = (int * lineq list) list; | 
| 308 | datatype result = Success of injust | Failure of history; | |
| 309 | ||
| 44654 | 310 | fun elim ctxt (ineqs, hist) = | 
| 311 | let val _ = print_ineqs ctxt ineqs | |
| 20217 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19618diff
changeset | 312 | val (triv, nontriv) = List.partition is_trivial ineqs in | 
| 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19618diff
changeset | 313 | if not (null triv) | 
| 59584 | 314 | then case find_first is_contradictory triv of | 
| 44654 | 315 | NONE => elim ctxt (nontriv, hist) | 
| 15531 | 316 | | SOME(Lineq(_,_,_,j)) => Success j | 
| 5982 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 317 | else | 
| 20217 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19618diff
changeset | 318 | if null nontriv then Failure hist | 
| 13498 | 319 | else | 
| 20217 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19618diff
changeset | 320 | let val (eqs, noneqs) = List.partition (fn (Lineq(_,ty,_,_)) => ty=Eq) nontriv in | 
| 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19618diff
changeset | 321 | if not (null eqs) then | 
| 31510 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 322 | let val c = | 
| 33042 | 323 | fold (fn Lineq(_,_,l,_) => fn cs => union (op =) l cs) eqs [] | 
| 31510 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 324 | |> filter (fn i => i <> 0) | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58839diff
changeset | 325 | |> sort (int_ord o apply2 abs) | 
| 31510 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 326 | |> hd | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 327 | val (eq as Lineq(_,_,ceq,_),othereqs) = | 
| 36692 
54b64d4ad524
farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
 haftmann parents: 
35872diff
changeset | 328 | extract_first (fn Lineq(_,_,l,_) => member (op =) l c) eqs | 
| 31986 | 329 | val v = find_index (fn v => v = c) ceq | 
| 23063 | 330 | val (ioth,roth) = List.partition (fn (Lineq(_,_,l,_)) => nth l v = 0) | 
| 5982 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 331 | (othereqs @ noneqs) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 332 | val others = map (elim_var v eq) roth @ ioth | 
| 44654 | 333 | in elim ctxt (others,(v,nontriv)::hist) end | 
| 5982 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 334 | else | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 335 | let val lists = map (fn (Lineq(_,_,l,_)) => l) noneqs | 
| 23063 | 336 | val numlist = 0 upto (length (hd lists) - 1) | 
| 337 | val coeffs = map (fn i => map (fn xs => nth xs i) lists) numlist | |
| 5982 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 338 | val blows = map calc_blowup coeffs | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 339 | val iblows = blows ~~ numlist | 
| 23063 | 340 | val nziblows = filter_out (fn (i, _) => i = 0) iblows | 
| 13498 | 341 | in if null nziblows then Failure((~1,nontriv)::hist) | 
| 342 | else | |
| 60348 | 343 | let val (_,v) = hd(sort (fn (x,y) => int_ord(fst(x),fst(y))) nziblows) | 
| 23063 | 344 | val (no,yes) = List.partition (fn (Lineq(_,_,l,_)) => nth l v = 0) ineqs | 
| 345 | val (pos,neg) = List.partition(fn (Lineq(_,_,l,_)) => nth l v > 0) yes | |
| 44654 | 346 | in elim ctxt (no @ map_product (elim_var v) pos neg, (v,nontriv)::hist) end | 
| 5982 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 347 | end | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 348 | end | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 349 | end; | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 350 | |
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 351 | (* ------------------------------------------------------------------------- *) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 352 | (* Translate back a proof. *) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 353 | (* ------------------------------------------------------------------------- *) | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 354 | |
| 44654 | 355 | fun trace_thm ctxt msgs th = | 
| 356 | (if Config.get ctxt LA_Data.trace | |
| 61268 | 357 | then tracing (cat_lines (msgs @ [Thm.string_of_thm ctxt th])) | 
| 44654 | 358 | else (); th); | 
| 9073 | 359 | |
| 44654 | 360 | fun trace_term ctxt msgs t = | 
| 361 | (if Config.get ctxt LA_Data.trace | |
| 362 | then tracing (cat_lines (msgs @ [Syntax.string_of_term ctxt t])) | |
| 363 | else (); t); | |
| 24076 | 364 | |
| 44654 | 365 | fun trace_msg ctxt msg = | 
| 366 | if Config.get ctxt LA_Data.trace then tracing msg else (); | |
| 9073 | 367 | |
| 52131 | 368 | val union_term = union Envir.aeconv; | 
| 26835 
404550067389
Lookup and union operations on terms are now modulo eta conversion.
 berghofe parents: 
24920diff
changeset | 369 | |
| 31510 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 370 | fun add_atoms (lhs, _, _, rhs, _, _) = | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 371 | union_term (map fst lhs) o union_term (map fst rhs); | 
| 6056 | 372 | |
| 31510 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 373 | fun atoms_of ds = fold add_atoms ds []; | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 374 | |
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 375 | (* | 
| 6056 | 376 | Simplification may detect a contradiction 'prematurely' due to type | 
| 377 | information: n+1 <= 0 is simplified to False and does not need to be crossed | |
| 378 | with 0 <= n. | |
| 379 | *) | |
| 380 | local | |
| 66035 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 381 | exception FalseE of thm * (int * cterm) list * Proof.context | 
| 6056 | 382 | in | 
| 27020 
b5b8afc9fdcd
added warning_count for issued reconstruction failure messages (limit 10);
 wenzelm parents: 
26945diff
changeset | 383 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 384 | fun mkthm ctxt asms (just: injust) = | 
| 24076 | 385 | let | 
| 42361 | 386 | val thy = Proof_Context.theory_of ctxt; | 
| 61097 | 387 |     val {add_mono_thms = add_mono_thms0, mult_mono_thms = mult_mono_thms0,
 | 
| 388 | inj_thms = inj_thms0, lessD = lessD0, simpset, ...} = get_data ctxt; | |
| 389 | val add_mono_thms = map (Thm.transfer thy) add_mono_thms0; | |
| 390 | val mult_mono_thms = map (Thm.transfer thy) mult_mono_thms0; | |
| 391 | val inj_thms = map (Thm.transfer thy) inj_thms0; | |
| 392 | val lessD = map (Thm.transfer thy) lessD0; | |
| 393 | ||
| 38763 
283f1f9969ba
Fast_Lin_Arith.number_of: more conventional merge that prefers the left side -- note that former ordering wrt. serial numbers makes it depend on accidental load order;
 wenzelm parents: 
38762diff
changeset | 394 | val number_of = number_of ctxt; | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 395 | val simpset_ctxt = put_simpset simpset ctxt; | 
| 31510 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 396 | fun only_concl f thm = | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 397 | if Thm.no_prems thm then f (Thm.concl_of thm) else NONE; | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 398 | val atoms = atoms_of (map_filter (only_concl (LA_Data.decomp ctxt)) asms); | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 399 | |
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 400 | fun use_first rules thm = | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 401 | get_first (fn th => SOME (thm RS th) handle THM _ => NONE) rules | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 402 | |
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 403 | fun add2 thm1 thm2 = | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 404 | use_first add_mono_thms (thm1 RS (thm2 RS LA_Logic.conjI)); | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 405 | fun try_add thms thm = get_first (fn th => add2 th thm) thms; | 
| 6056 | 406 | |
| 31510 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 407 | fun add_thms thm1 thm2 = | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 408 | (case add2 thm1 thm2 of | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 409 | NONE => | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 410 | (case try_add ([thm1] RL inj_thms) thm2 of | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 411 | NONE => | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 412 | (the (try_add ([thm2] RL inj_thms) thm1) | 
| 51930 
52fd62618631
prefer explicitly qualified exceptions, which is particular important for robust handlers;
 wenzelm parents: 
51717diff
changeset | 413 | handle Option.Option => | 
| 44654 | 414 | (trace_thm ctxt [] thm1; trace_thm ctxt [] thm2; | 
| 40316 
665862241968
replaced ancient sys_error by raise Fail, assuming that the latter is not handled specifically by the environment;
 wenzelm parents: 
38763diff
changeset | 415 | raise Fail "Linear arithmetic: failed to add thms")) | 
| 31510 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 416 | | SOME thm => thm) | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 417 | | SOME thm => thm); | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 418 | |
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 419 | fun mult_by_add n thm = | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 420 | let fun mul i th = if i = 1 then th else mul (i - 1) (add_thms thm th) | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 421 | in mul n thm end; | 
| 10575 | 422 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 423 | val rewr = Simplifier.rewrite simpset_ctxt; | 
| 31510 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 424 | val rewrite_concl = Conv.fconv_rule (Conv.concl_conv ~1 (Conv.arg_conv | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 425 | (Conv.binop_conv rewr))); | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 426 | fun discharge_prem thm = if Thm.nprems_of thm = 0 then thm else | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 427 | let val cv = Conv.arg1_conv (Conv.arg_conv rewr) | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 428 | in Thm.implies_elim (Conv.fconv_rule cv thm) LA_Logic.trueI end | 
| 20217 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19618diff
changeset | 429 | |
| 31510 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 430 | fun mult n thm = | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 431 | (case use_first mult_mono_thms thm of | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 432 | NONE => mult_by_add n thm | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 433 | | SOME mth => | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 434 | let | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 435 | val cv = mth |> Thm.cprop_of |> Drule.strip_imp_concl | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 436 | |> Thm.dest_arg |> Thm.dest_arg1 |> Thm.dest_arg1 | 
| 59586 | 437 | val T = Thm.typ_of_cterm cv | 
| 31510 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 438 | in | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 439 | mth | 
| 77879 | 440 | |> Thm.instantiate (TVars.empty, Vars.make1 (dest_Var (Thm.term_of cv), number_of T n)) | 
| 31510 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 441 | |> rewrite_concl | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 442 | |> discharge_prem | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 443 | handle CTERM _ => mult_by_add n thm | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 444 | | THM _ => mult_by_add n thm | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 445 | end); | 
| 10691 | 446 | |
| 66035 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 447 | fun mult_thm n thm = | 
| 31510 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 448 | if n = ~1 then thm RS LA_Logic.sym | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 449 | else if n < 0 then mult (~n) thm RS LA_Logic.sym | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 450 | else mult n thm; | 
| 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 451 | |
| 66035 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 452 | fun simp thm (cx as (_, hyps, ctxt')) = | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 453 | let val thm' = trace_thm ctxt ["Simplified:"] (full_simplify simpset_ctxt thm) | 
| 66035 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 454 | in if LA_Logic.is_False thm' then raise FalseE (thm', hyps, ctxt') else (thm', cx) end; | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 455 | |
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 456 | fun abs_thm i (cx as (terms, hyps, ctxt)) = | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 457 | (case AList.lookup (op =) hyps i of | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 458 | SOME ct => (Thm.assume ct, cx) | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 459 | | NONE => | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 460 | let | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 461 | val thm = nth asms i | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 462 | val (t, (terms', ctxt')) = LA_Data.abstract (Thm.prop_of thm) (terms, ctxt) | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 463 | val ct = Thm.cterm_of ctxt' t | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 464 | in (Thm.assume ct, (terms', (i, ct) :: hyps, ctxt')) end); | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 465 | |
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 466 | fun nat_thm t (terms, hyps, ctxt) = | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 467 | let val (t', (terms', ctxt')) = LA_Data.abstract_arith t (terms, ctxt) | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 468 | in (LA_Logic.mk_nat_thm thy t', (terms', hyps, ctxt')) end; | 
| 6056 | 469 | |
| 66035 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 470 | fun step0 msg (thm, cx) = (trace_thm ctxt [msg] thm, cx) | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 471 | fun step1 msg j f cx = mk j cx |>> f |>> trace_thm ctxt [msg] | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 472 | and step2 msg j1 j2 f cx = mk j1 cx ||>> mk j2 |>> f |>> trace_thm ctxt [msg] | 
| 5982 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 473 | |
| 66035 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 474 |     and mk (Asm i) cx = step0 ("Asm " ^ string_of_int i) (abs_thm i cx)
 | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 475 |       | mk (Nat i) cx = step0 ("Nat " ^ string_of_int i) (nat_thm (nth atoms i) cx)
 | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 476 | | mk (LessD j) cx = step1 "L" j (fn thm => hd ([thm] RL lessD)) cx | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 477 | | mk (NotLeD j) cx = step1 "NLe" j (fn thm => thm RS LA_Logic.not_leD) cx | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 478 | | mk (NotLeDD j) cx = step1 "NLeD" j (fn thm => hd ([thm RS LA_Logic.not_leD] RL lessD)) cx | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 479 | | mk (NotLessD j) cx = step1 "NL" j (fn thm => thm RS LA_Logic.not_lessD) cx | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 480 | | mk (Added (j1, j2)) cx = step2 "+" j1 j2 (uncurry add_thms) cx |-> simp | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 481 | | mk (Multiplied (n, j)) cx = | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 482 |           (trace_msg ctxt ("*" ^ string_of_int n); step1 "*" j (mult_thm n) cx)
 | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 483 | |
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 484 | fun finish ctxt' hyps thm = | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 485 | thm | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 486 | |> fold_rev (Thm.implies_intr o snd) hyps | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 487 | |> singleton (Variable.export ctxt' ctxt) | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 488 | |> fold (fn (i, _) => fn thm => nth asms i RS thm) hyps | 
| 27020 
b5b8afc9fdcd
added warning_count for issued reconstruction failure messages (limit 10);
 wenzelm parents: 
26945diff
changeset | 489 | in | 
| 
b5b8afc9fdcd
added warning_count for issued reconstruction failure messages (limit 10);
 wenzelm parents: 
26945diff
changeset | 490 | let | 
| 44654 | 491 | val _ = trace_msg ctxt "mkthm"; | 
| 66035 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 492 | val (thm, (_, hyps, ctxt')) = mk just ([], [], ctxt); | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 493 | val _ = trace_thm ctxt ["Final thm:"] thm; | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 494 | val fls = simplify simpset_ctxt thm; | 
| 44654 | 495 | val _ = trace_thm ctxt ["After simplification:"] fls; | 
| 27020 
b5b8afc9fdcd
added warning_count for issued reconstruction failure messages (limit 10);
 wenzelm parents: 
26945diff
changeset | 496 | val _ = | 
| 
b5b8afc9fdcd
added warning_count for issued reconstruction failure messages (limit 10);
 wenzelm parents: 
26945diff
changeset | 497 | if LA_Logic.is_False fls then () | 
| 
b5b8afc9fdcd
added warning_count for issued reconstruction failure messages (limit 10);
 wenzelm parents: 
26945diff
changeset | 498 | else | 
| 35872 
9b579860d59b
removed warning_count (known causes for warnings have been resolved)
 boehmes parents: 
35861diff
changeset | 499 | (tracing (cat_lines | 
| 61268 | 500 | (["Assumptions:"] @ map (Thm.string_of_thm ctxt) asms @ [""] @ | 
| 501 | ["Proved:", Thm.string_of_thm ctxt fls, ""])); | |
| 78133 | 502 | warning "Linear arithmetic should have refuted the assumptions but failed to.") | 
| 66035 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 503 | in finish ctxt' hyps fls end | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 504 | handle FalseE (thm, hyps, ctxt') => | 
| 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 505 | trace_thm ctxt ["False reached early:"] (finish ctxt' hyps thm) | 
| 27020 
b5b8afc9fdcd
added warning_count for issued reconstruction failure messages (limit 10);
 wenzelm parents: 
26945diff
changeset | 506 | end; | 
| 
b5b8afc9fdcd
added warning_count for issued reconstruction failure messages (limit 10);
 wenzelm parents: 
26945diff
changeset | 507 | |
| 6056 | 508 | end; | 
| 5982 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 509 | |
| 23261 | 510 | fun coeff poly atom = | 
| 52131 | 511 | AList.lookup Envir.aeconv poly atom |> the_default 0; | 
| 10691 | 512 | |
| 513 | fun integ(rlhs,r,rel,rrhs,s,d) = | |
| 63201 | 514 | let val (rn,rd) = Rat.dest r and (sn,sd) = Rat.dest s | 
| 63227 
d3ed7f00e818
Integer.lcm normalizes the sign as in HOL/GCD.thy;
 wenzelm parents: 
63201diff
changeset | 515 | val m = Integer.lcms(map (snd o Rat.dest) (r :: s :: map snd rlhs @ map snd rrhs)) | 
| 22846 | 516 | fun mult(t,r) = | 
| 63201 | 517 | let val (i,j) = Rat.dest r | 
| 15965 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 paulson parents: 
15922diff
changeset | 518 | in (t,i * (m div j)) end | 
| 12932 | 519 | in (m,(map mult rlhs, rn*(m div rd), rel, map mult rrhs, sn*(m div sd), d)) end | 
| 10691 | 520 | |
| 38052 | 521 | fun mklineq atoms = | 
| 20217 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19618diff
changeset | 522 | fn (item, k) => | 
| 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19618diff
changeset | 523 | let val (m, (lhs,i,rel,rhs,j,discrete)) = integ item | 
| 13498 | 524 | val lhsa = map (coeff lhs) atoms | 
| 525 | and rhsa = map (coeff rhs) atoms | |
| 18330 | 526 | val diff = map2 (curry (op -)) rhsa lhsa | 
| 13498 | 527 | val c = i-j | 
| 528 | val just = Asm k | |
| 31511 | 529 | fun lineq(c,le,cs,j) = Lineq(c,le,cs, if m=1 then j else Multiplied(m,j)) | 
| 13498 | 530 | in case rel of | 
| 531 | "<=" => lineq(c,Le,diff,just) | |
| 532 | | "~<=" => if discrete | |
| 533 | then lineq(1-c,Le,map (op ~) diff,NotLeDD(just)) | |
| 534 | else lineq(~c,Lt,map (op ~) diff,NotLeD(just)) | |
| 535 | | "<" => if discrete | |
| 536 | then lineq(c+1,Le,diff,LessD(just)) | |
| 537 | else lineq(c,Lt,diff,just) | |
| 538 | | "~<" => lineq(~c,Le,map (op~) diff,NotLessD(just)) | |
| 539 | | "=" => lineq(c,Eq,diff,just) | |
| 40316 
665862241968
replaced ancient sys_error by raise Fail, assuming that the latter is not handled specifically by the environment;
 wenzelm parents: 
38763diff
changeset | 540 |      | _     => raise Fail ("mklineq" ^ rel)
 | 
| 5982 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 541 | end; | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 542 | |
| 13498 | 543 | (* ------------------------------------------------------------------------- *) | 
| 544 | (* Print (counter) example *) | |
| 545 | (* ------------------------------------------------------------------------- *) | |
| 546 | ||
| 20217 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19618diff
changeset | 547 | (* ------------------------------------------------------------------------- *) | 
| 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19618diff
changeset | 548 | |
| 20268 | 549 | fun mknat (pTs : typ list) (ixs : int list) (atom : term, i : int) : lineq option = | 
| 20217 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19618diff
changeset | 550 | if LA_Logic.is_nat (pTs, atom) | 
| 6056 | 551 | then let val l = map (fn j => if j=i then 1 else 0) ixs | 
| 20217 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19618diff
changeset | 552 | in SOME (Lineq (0, Le, l, Nat i)) end | 
| 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19618diff
changeset | 553 | else NONE; | 
| 6056 | 554 | |
| 13186 
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
 nipkow parents: 
13105diff
changeset | 555 | (* This code is tricky. It takes a list of premises in the order they occur | 
| 15531 | 556 | in the subgoal. Numerical premises are coded as SOME(tuple), non-numerical | 
| 557 | ones as NONE. Going through the premises, each numeric one is converted into | |
| 13186 
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
 nipkow parents: 
13105diff
changeset | 558 | a Lineq. The tricky bit is to convert ~= which is split into two cases < and | 
| 13498 | 559 | >. Thus split_items returns a list of equation systems. This may blow up if | 
| 13186 
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
 nipkow parents: 
13105diff
changeset | 560 | there are many ~=, but in practice it does not seem to happen. The really | 
| 
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
 nipkow parents: 
13105diff
changeset | 561 | tricky bit is to arrange the order of the cases such that they coincide with | 
| 
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
 nipkow parents: 
13105diff
changeset | 562 | the order in which the cases are in the end generated by the tactic that | 
| 
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
 nipkow parents: 
13105diff
changeset | 563 | applies the generated refutation thms (see function 'refute_tac'). | 
| 
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
 nipkow parents: 
13105diff
changeset | 564 | |
| 
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
 nipkow parents: 
13105diff
changeset | 565 | For variables n of type nat, a constraint 0 <= n is added. | 
| 
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
 nipkow parents: 
13105diff
changeset | 566 | *) | 
| 20217 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19618diff
changeset | 567 | |
| 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19618diff
changeset | 568 | (* FIXME: To optimize, the splitting of cases and the search for refutations *) | 
| 20276 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 569 | (* could be intertwined: separate the first (fully split) case, *) | 
| 20217 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19618diff
changeset | 570 | (* refute it, continue with splitting and refuting. Terminate with *) | 
| 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19618diff
changeset | 571 | (* failure as soon as a case could not be refuted; i.e. delay further *) | 
| 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19618diff
changeset | 572 | (* splitting until after a refutation for other cases has been found. *) | 
| 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19618diff
changeset | 573 | |
| 30406 
15dc25f8a0e2
Instead of giving up entirely, arith now ignores all inequalities when there are too many.
 webertj parents: 
27020diff
changeset | 574 | fun split_items ctxt do_pre split_neq (Ts, terms) : (typ list * (LA_Data.decomp * int) list) list = | 
| 20276 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 575 | let | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 576 | (* splits inequalities '~=' into '<' and '>'; this corresponds to *) | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 577 | (* 'REPEAT_DETERM (eresolve_tac neqE i)' at the theorem/tactic *) | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 578 | (* level *) | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 579 | (* FIXME: this is currently sensitive to the order of theorems in *) | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 580 | (* neqE: The theorem for type "nat" must come first. A *) | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 581 | (* better (i.e. less likely to break when neqE changes) *) | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 582 | (* implementation should *test* which theorem from neqE *) | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 583 | (* can be applied, and split the premise accordingly. *) | 
| 26945 | 584 | fun elim_neq (ineqs : (LA_Data.decomp option * bool) list) : | 
| 585 | (LA_Data.decomp option * bool) list list = | |
| 20276 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 586 | let | 
| 66035 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 boehmes parents: 
63227diff
changeset | 587 | fun elim_neq' _ ([] : (LA_Data.decomp option * bool) list) : | 
| 26945 | 588 | (LA_Data.decomp option * bool) list list = | 
| 20276 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 589 | [[]] | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 590 | | elim_neq' nat_only ((NONE, is_nat) :: ineqs) = | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 591 | map (cons (NONE, is_nat)) (elim_neq' nat_only ineqs) | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 592 | | elim_neq' nat_only ((ineq as (SOME (l, i, rel, r, j, d), is_nat)) :: ineqs) = | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 593 | if rel = "~=" andalso (not nat_only orelse is_nat) then | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 594 | (* [| ?l ~= ?r; ?l < ?r ==> ?R; ?r < ?l ==> ?R |] ==> ?R *) | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 595 | elim_neq' nat_only (ineqs @ [(SOME (l, i, "<", r, j, d), is_nat)]) @ | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 596 | elim_neq' nat_only (ineqs @ [(SOME (r, j, "<", l, i, d), is_nat)]) | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 597 | else | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 598 | map (cons ineq) (elim_neq' nat_only ineqs) | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 599 | in | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 600 | ineqs |> elim_neq' true | 
| 26945 | 601 | |> maps (elim_neq' false) | 
| 20276 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 602 | end | 
| 13464 | 603 | |
| 30406 
15dc25f8a0e2
Instead of giving up entirely, arith now ignores all inequalities when there are too many.
 webertj parents: 
27020diff
changeset | 604 | fun ignore_neq (NONE, bool) = (NONE, bool) | 
| 
15dc25f8a0e2
Instead of giving up entirely, arith now ignores all inequalities when there are too many.
 webertj parents: 
27020diff
changeset | 605 | | ignore_neq (ineq as SOME (_, _, rel, _, _, _), bool) = | 
| 
15dc25f8a0e2
Instead of giving up entirely, arith now ignores all inequalities when there are too many.
 webertj parents: 
27020diff
changeset | 606 | if rel = "~=" then (NONE, bool) else (ineq, bool) | 
| 
15dc25f8a0e2
Instead of giving up entirely, arith now ignores all inequalities when there are too many.
 webertj parents: 
27020diff
changeset | 607 | |
| 20276 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 608 | fun number_hyps _ [] = [] | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 609 | | number_hyps n (NONE::xs) = number_hyps (n+1) xs | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 610 | | number_hyps n ((SOME x)::xs) = (x, n) :: number_hyps (n+1) xs | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 611 | |
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 612 | val result = (Ts, terms) | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 613 | |> (* user-defined preprocessing of the subgoal *) | 
| 24076 | 614 | (if do_pre then LA_Data.pre_decomp ctxt else Library.single) | 
| 44654 | 615 |     |> tap (fn subgoals => trace_msg ctxt ("Preprocessing yields " ^
 | 
| 23195 | 616 | string_of_int (length subgoals) ^ " subgoal(s) total.")) | 
| 22846 | 617 | |> (* produce the internal encoding of (in-)equalities *) | 
| 24076 | 618 | map (apsnd (map (fn t => (LA_Data.decomp ctxt t, LA_Data.domain_is_nat t)))) | 
| 20276 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 619 | |> (* splitting of inequalities *) | 
| 30406 
15dc25f8a0e2
Instead of giving up entirely, arith now ignores all inequalities when there are too many.
 webertj parents: 
27020diff
changeset | 620 | map (apsnd (if split_neq then elim_neq else | 
| 
15dc25f8a0e2
Instead of giving up entirely, arith now ignores all inequalities when there are too many.
 webertj parents: 
27020diff
changeset | 621 | Library.single o map ignore_neq)) | 
| 22846 | 622 | |> maps (fn (Ts, subgoals) => map (pair Ts o map fst) subgoals) | 
| 20276 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 623 | |> (* numbering of hypotheses, ignoring irrelevant ones *) | 
| 
d94dc40673b1
possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
 webertj parents: 
20268diff
changeset | 624 | map (apsnd (number_hyps 0)) | 
| 23195 | 625 | in | 
| 44654 | 626 |   trace_msg ctxt ("Splitting of inequalities yields " ^
 | 
| 23195 | 627 | string_of_int (length result) ^ " subgoal(s) total."); | 
| 628 | result | |
| 629 | end; | |
| 13464 | 630 | |
| 59656 | 631 | fun refutes ctxt : | 
| 26945 | 632 | (typ list * (LA_Data.decomp * int) list) list -> injust list -> injust list option = | 
| 633 | let | |
| 634 | fun refute ((Ts, initems : (LA_Data.decomp * int) list) :: initemss) (js: injust list) = | |
| 635 | let | |
| 31510 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 boehmes parents: 
31102diff
changeset | 636 | val atoms = atoms_of (map fst initems) | 
| 26945 | 637 | val n = length atoms | 
| 38052 | 638 | val mkleq = mklineq atoms | 
| 26945 | 639 | val ixs = 0 upto (n - 1) | 
| 640 | val iatoms = atoms ~~ ixs | |
| 32952 | 641 | val natlineqs = map_filter (mknat Ts ixs) iatoms | 
| 26945 | 642 | val ineqs = map mkleq initems @ natlineqs | 
| 59656 | 643 | in | 
| 644 | (case elim ctxt (ineqs, []) of | |
| 26945 | 645 | Success j => | 
| 44654 | 646 |                  (trace_msg ctxt ("Contradiction! (" ^ string_of_int (length js + 1) ^ ")");
 | 
| 26945 | 647 | refute initemss (js @ [j])) | 
| 59656 | 648 | | Failure _ => NONE) | 
| 26945 | 649 | end | 
| 650 | | refute [] js = SOME js | |
| 651 | in refute end; | |
| 5982 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 652 | |
| 59656 | 653 | fun refute ctxt params do_pre split_neq terms : injust list option = | 
| 654 | refutes ctxt (split_items ctxt do_pre split_neq (map snd params, terms)) []; | |
| 20254 
58b71535ed00
lin_arith_prover splits certain operators (e.g. min, max, abs)
 webertj parents: 
20217diff
changeset | 655 | |
| 22950 | 656 | fun count P xs = length (filter P xs); | 
| 20254 
58b71535ed00
lin_arith_prover splits certain operators (e.g. min, max, abs)
 webertj parents: 
20217diff
changeset | 657 | |
| 59656 | 658 | fun prove ctxt params do_pre Hs concl : bool * injust list option = | 
| 20254 
58b71535ed00
lin_arith_prover splits certain operators (e.g. min, max, abs)
 webertj parents: 
20217diff
changeset | 659 | let | 
| 44654 | 660 | val _ = trace_msg ctxt "prove:" | 
| 20254 
58b71535ed00
lin_arith_prover splits certain operators (e.g. min, max, abs)
 webertj parents: 
20217diff
changeset | 661 | (* append the negated conclusion to 'Hs' -- this corresponds to *) | 
| 
58b71535ed00
lin_arith_prover splits certain operators (e.g. min, max, abs)
 webertj parents: 
20217diff
changeset | 662 | (* 'DETERM (resolve_tac [LA_Logic.notI, LA_Logic.ccontr] i)' at the *) | 
| 
58b71535ed00
lin_arith_prover splits certain operators (e.g. min, max, abs)
 webertj parents: 
20217diff
changeset | 663 | (* theorem/tactic level *) | 
| 
58b71535ed00
lin_arith_prover splits certain operators (e.g. min, max, abs)
 webertj parents: 
20217diff
changeset | 664 | val Hs' = Hs @ [LA_Logic.neg_prop concl] | 
| 
58b71535ed00
lin_arith_prover splits certain operators (e.g. min, max, abs)
 webertj parents: 
20217diff
changeset | 665 | fun is_neq NONE = false | 
| 
58b71535ed00
lin_arith_prover splits certain operators (e.g. min, max, abs)
 webertj parents: 
20217diff
changeset | 666 | | is_neq (SOME (_,_,r,_,_,_)) = (r = "~=") | 
| 44654 | 667 | val neq_limit = Config.get ctxt LA_Data.neq_limit | 
| 30406 
15dc25f8a0e2
Instead of giving up entirely, arith now ignores all inequalities when there are too many.
 webertj parents: 
27020diff
changeset | 668 | val split_neq = count is_neq (map (LA_Data.decomp ctxt) Hs') <= neq_limit | 
| 20254 
58b71535ed00
lin_arith_prover splits certain operators (e.g. min, max, abs)
 webertj parents: 
20217diff
changeset | 669 | in | 
| 30406 
15dc25f8a0e2
Instead of giving up entirely, arith now ignores all inequalities when there are too many.
 webertj parents: 
27020diff
changeset | 670 | if split_neq then () | 
| 24076 | 671 | else | 
| 44654 | 672 |       trace_msg ctxt ("neq_limit exceeded (current value is " ^
 | 
| 30406 
15dc25f8a0e2
Instead of giving up entirely, arith now ignores all inequalities when there are too many.
 webertj parents: 
27020diff
changeset | 673 | string_of_int neq_limit ^ "), ignoring all inequalities"); | 
| 59656 | 674 | (split_neq, refute ctxt params do_pre split_neq Hs') | 
| 23190 | 675 |   end handle TERM ("neg_prop", _) =>
 | 
| 676 | (* since no meta-logic negation is available, we can only fail if *) | |
| 677 | (* the conclusion is not of the form 'Trueprop $ _' (simply *) | |
| 678 | (* dropping the conclusion doesn't work either, because even *) | |
| 679 | (* 'False' does not imply arbitrary 'concl::prop') *) | |
| 44654 | 680 | (trace_msg ctxt "prove failed (cannot negate conclusion)."; | 
| 30406 
15dc25f8a0e2
Instead of giving up entirely, arith now ignores all inequalities when there are too many.
 webertj parents: 
27020diff
changeset | 681 | (false, NONE)); | 
| 20217 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19618diff
changeset | 682 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 683 | fun refute_tac ctxt (i, split_neq, justs) = | 
| 6074 | 684 | fn state => | 
| 24076 | 685 | let | 
| 32091 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
31986diff
changeset | 686 | val _ = trace_thm ctxt | 
| 44654 | 687 | ["refute_tac (on subgoal " ^ string_of_int i ^ ", with " ^ | 
| 688 | string_of_int (length justs) ^ " justification(s)):"] state | |
| 61097 | 689 | val neqE = get_neqE ctxt; | 
| 24076 | 690 | fun just1 j = | 
| 691 | (* eliminate inequalities *) | |
| 30406 
15dc25f8a0e2
Instead of giving up entirely, arith now ignores all inequalities when there are too many.
 webertj parents: 
27020diff
changeset | 692 | (if split_neq then | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59058diff
changeset | 693 | REPEAT_DETERM (eresolve_tac ctxt neqE i) | 
| 30406 
15dc25f8a0e2
Instead of giving up entirely, arith now ignores all inequalities when there are too many.
 webertj parents: 
27020diff
changeset | 694 | else | 
| 
15dc25f8a0e2
Instead of giving up entirely, arith now ignores all inequalities when there are too many.
 webertj parents: 
27020diff
changeset | 695 | all_tac) THEN | 
| 44654 | 696 | PRIMITIVE (trace_thm ctxt ["State after neqE:"]) THEN | 
| 24076 | 697 | (* use theorems generated from the actual justifications *) | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59058diff
changeset | 698 |           Subgoal.FOCUS (fn {prems, ...} => resolve_tac ctxt [mkthm ctxt prems j] 1) ctxt i
 | 
| 24076 | 699 | in | 
| 700 | (* rewrite "[| A1; ...; An |] ==> B" to "[| A1; ...; An; ~B |] ==> False" *) | |
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59058diff
changeset | 701 | DETERM (resolve_tac ctxt [LA_Logic.notI, LA_Logic.ccontr] i) THEN | 
| 24076 | 702 | (* user-defined preprocessing of the subgoal *) | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 703 | DETERM (LA_Data.pre_tac ctxt i) THEN | 
| 44654 | 704 | PRIMITIVE (trace_thm ctxt ["State after pre_tac:"]) THEN | 
| 24076 | 705 | (* prove every resulting subgoal, using its justification *) | 
| 706 | EVERY (map just1 justs) | |
| 20217 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19618diff
changeset | 707 | end state; | 
| 6074 | 708 | |
| 5982 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 709 | (* | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 710 | Fast but very incomplete decider. Only premises and conclusions | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 711 | that are already (negated) (in)equations are taken into account. | 
| 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 712 | *) | 
| 59656 | 713 | fun simpset_lin_arith_tac ctxt = SUBGOAL (fn (A, i) => | 
| 24076 | 714 | let | 
| 715 | val params = rev (Logic.strip_params A) | |
| 716 | val Hs = Logic.strip_assums_hyp A | |
| 717 | val concl = Logic.strip_assums_concl A | |
| 44654 | 718 | val _ = trace_term ctxt ["Trying to refute subgoal " ^ string_of_int i] A | 
| 24076 | 719 | in | 
| 59656 | 720 | case prove ctxt params true Hs concl of | 
| 44654 | 721 | (_, NONE) => (trace_msg ctxt "Refutation failed."; no_tac) | 
| 722 | | (split_neq, SOME js) => (trace_msg ctxt "Refutation succeeded."; | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 723 | refute_tac ctxt (i, split_neq, js)) | 
| 24076 | 724 | end); | 
| 5982 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 725 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 726 | fun prems_lin_arith_tac ctxt = | 
| 61841 
4d3527b94f2a
more general types Proof.method / context_tactic;
 wenzelm parents: 
61268diff
changeset | 727 | Method.insert_tac ctxt (Simplifier.prems_of ctxt) THEN' | 
| 59656 | 728 | simpset_lin_arith_tac ctxt; | 
| 17613 | 729 | |
| 24076 | 730 | fun lin_arith_tac ctxt = | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 731 | simpset_lin_arith_tac (empty_simpset ctxt); | 
| 24076 | 732 | |
| 733 | ||
| 5982 
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
 nipkow parents: diff
changeset | 734 | |
| 13186 
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
 nipkow parents: 
13105diff
changeset | 735 | (** Forward proof from theorems **) | 
| 
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
 nipkow parents: 
13105diff
changeset | 736 | |
| 20433 
55471f940e5c
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20280diff
changeset | 737 | (* More tricky code. Needs to arrange the proofs of the multiple cases (due | 
| 
55471f940e5c
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20280diff
changeset | 738 | to splits of ~= premises) such that it coincides with the order of the cases | 
| 
55471f940e5c
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20280diff
changeset | 739 | generated by function split_items. *) | 
| 
55471f940e5c
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20280diff
changeset | 740 | |
| 
55471f940e5c
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20280diff
changeset | 741 | datatype splittree = Tip of thm list | 
| 
55471f940e5c
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20280diff
changeset | 742 | | Spl of thm * cterm * splittree * cterm * splittree; | 
| 
55471f940e5c
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20280diff
changeset | 743 | |
| 
55471f940e5c
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20280diff
changeset | 744 | (* "(ct1 ==> ?R) ==> (ct2 ==> ?R) ==> ?R" is taken to (ct1, ct2) *) | 
| 
55471f940e5c
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20280diff
changeset | 745 | |
| 
55471f940e5c
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20280diff
changeset | 746 | fun extract (imp : cterm) : cterm * cterm = | 
| 
55471f940e5c
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20280diff
changeset | 747 | let val (Il, r) = Thm.dest_comb imp | 
| 
55471f940e5c
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20280diff
changeset | 748 | val (_, imp1) = Thm.dest_comb Il | 
| 
55471f940e5c
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20280diff
changeset | 749 | val (Ict1, _) = Thm.dest_comb imp1 | 
| 
55471f940e5c
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20280diff
changeset | 750 | val (_, ct1) = Thm.dest_comb Ict1 | 
| 
55471f940e5c
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20280diff
changeset | 751 | val (Ir, _) = Thm.dest_comb r | 
| 
55471f940e5c
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20280diff
changeset | 752 | val (_, Ict2r) = Thm.dest_comb Ir | 
| 
55471f940e5c
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20280diff
changeset | 753 | val (Ict2, _) = Thm.dest_comb Ict2r | 
| 
55471f940e5c
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20280diff
changeset | 754 | val (_, ct2) = Thm.dest_comb Ict2 | 
| 
55471f940e5c
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20280diff
changeset | 755 | in (ct1, ct2) end; | 
| 
55471f940e5c
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20280diff
changeset | 756 | |
| 24076 | 757 | fun splitasms ctxt (asms : thm list) : splittree = | 
| 61097 | 758 | let val neqE = get_neqE ctxt | 
| 35693 
d58a4ac1ca1c
Use same order of neq-elimination as in proof search.
 hoelzl parents: 
35230diff
changeset | 759 | fun elim_neq [] (asms', []) = Tip (rev asms') | 
| 
d58a4ac1ca1c
Use same order of neq-elimination as in proof search.
 hoelzl parents: 
35230diff
changeset | 760 | | elim_neq [] (asms', asms) = Tip (rev asms' @ asms) | 
| 49387 | 761 | | elim_neq (_ :: neqs) (asms', []) = elim_neq neqs ([],rev asms') | 
| 35693 
d58a4ac1ca1c
Use same order of neq-elimination as in proof search.
 hoelzl parents: 
35230diff
changeset | 762 | | elim_neq (neqs as (neq :: _)) (asms', asm::asms) = | 
| 
d58a4ac1ca1c
Use same order of neq-elimination as in proof search.
 hoelzl parents: 
35230diff
changeset | 763 | (case get_first (fn th => SOME (asm COMP th) handle THM _ => NONE) [neq] of | 
| 20433 
55471f940e5c
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20280diff
changeset | 764 | SOME spl => | 
| 59582 | 765 | let val (ct1, ct2) = extract (Thm.cprop_of spl) | 
| 36945 | 766 | val thm1 = Thm.assume ct1 | 
| 767 | val thm2 = Thm.assume ct2 | |
| 35693 
d58a4ac1ca1c
Use same order of neq-elimination as in proof search.
 hoelzl parents: 
35230diff
changeset | 768 | in Spl (spl, ct1, elim_neq neqs (asms', asms@[thm1]), | 
| 
d58a4ac1ca1c
Use same order of neq-elimination as in proof search.
 hoelzl parents: 
35230diff
changeset | 769 | ct2, elim_neq neqs (asms', asms@[thm2])) | 
| 20433 
55471f940e5c
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20280diff
changeset | 770 | end | 
| 35693 
d58a4ac1ca1c
Use same order of neq-elimination as in proof search.
 hoelzl parents: 
35230diff
changeset | 771 | | NONE => elim_neq neqs (asm::asms', asms)) | 
| 
d58a4ac1ca1c
Use same order of neq-elimination as in proof search.
 hoelzl parents: 
35230diff
changeset | 772 | in elim_neq neqE ([], asms) end; | 
| 20433 
55471f940e5c
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20280diff
changeset | 773 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 774 | fun fwdproof ctxt (Tip asms : splittree) (j::js : injust list) = (mkthm ctxt asms j, js) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 775 | | fwdproof ctxt (Spl (thm, ct1, tree1, ct2, tree2)) js = | 
| 24076 | 776 | let | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 777 | val (thm1, js1) = fwdproof ctxt tree1 js | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 778 | val (thm2, js2) = fwdproof ctxt tree2 js1 | 
| 36945 | 779 | val thm1' = Thm.implies_intr ct1 thm1 | 
| 780 | val thm2' = Thm.implies_intr ct2 thm2 | |
| 24076 | 781 | in (thm2' COMP (thm1' COMP thm), js2) end; | 
| 782 | (* FIXME needs handle THM _ => NONE ? *) | |
| 20433 
55471f940e5c
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20280diff
changeset | 783 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 784 | fun prover ctxt thms Tconcl (js : injust list) split_neq pos : thm option = | 
| 24076 | 785 | let | 
| 786 | val nTconcl = LA_Logic.neg_prop Tconcl | |
| 59642 | 787 | val cnTconcl = Thm.cterm_of ctxt nTconcl | 
| 36945 | 788 | val nTconclthm = Thm.assume cnTconcl | 
| 30406 
15dc25f8a0e2
Instead of giving up entirely, arith now ignores all inequalities when there are too many.
 webertj parents: 
27020diff
changeset | 789 | val tree = (if split_neq then splitasms ctxt else Tip) (thms @ [nTconclthm]) | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 790 | val (Falsethm, _) = fwdproof ctxt tree js | 
| 24076 | 791 | val contr = if pos then LA_Logic.ccontr else LA_Logic.notI | 
| 36945 | 792 | val concl = Thm.implies_intr cnTconcl Falsethm COMP contr | 
| 44654 | 793 | in SOME (trace_thm ctxt ["Proved by lin. arith. prover:"] (LA_Logic.mk_Eq concl)) end | 
| 24076 | 794 | (*in case concl contains ?-var, which makes assume fail:*) (* FIXME Variable.import_terms *) | 
| 795 | handle THM _ => NONE; | |
| 13186 
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
 nipkow parents: 
13105diff
changeset | 796 | |
| 
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
 nipkow parents: 
13105diff
changeset | 797 | (* PRE: concl is not negated! | 
| 
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
 nipkow parents: 
13105diff
changeset | 798 | This assumption is OK because | 
| 24076 | 799 | 1. lin_arith_simproc tries both to prove and disprove concl and | 
| 800 | 2. lin_arith_simproc is applied by the Simplifier which | |
| 13186 
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
 nipkow parents: 
13105diff
changeset | 801 | dives into terms and will thus try the non-negated concl anyway. | 
| 
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
 nipkow parents: 
13105diff
changeset | 802 | *) | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 803 | fun lin_arith_simproc ctxt concl = | 
| 24076 | 804 | let | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 805 | val thms = maps LA_Logic.atomize (Simplifier.prems_of ctxt) | 
| 24076 | 806 | val Hs = map Thm.prop_of thms | 
| 61144 | 807 | val Tconcl = LA_Logic.mk_Trueprop (Thm.term_of concl) | 
| 24076 | 808 | in | 
| 59656 | 809 | case prove ctxt [] false Hs Tconcl of (* concl provable? *) | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 810 | (split_neq, SOME js) => prover ctxt thms Tconcl js split_neq true | 
| 30406 
15dc25f8a0e2
Instead of giving up entirely, arith now ignores all inequalities when there are too many.
 webertj parents: 
27020diff
changeset | 811 | | (_, NONE) => | 
| 24076 | 812 | let val nTconcl = LA_Logic.neg_prop Tconcl in | 
| 59656 | 813 | case prove ctxt [] false Hs nTconcl of (* ~concl provable? *) | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49387diff
changeset | 814 | (split_neq, SOME js) => prover ctxt thms nTconcl js split_neq false | 
| 30406 
15dc25f8a0e2
Instead of giving up entirely, arith now ignores all inequalities when there are too many.
 webertj parents: 
27020diff
changeset | 815 | | (_, NONE) => NONE | 
| 24076 | 816 | end | 
| 817 | end; | |
| 6074 | 818 | |
| 819 | end; |