src/Provers/Arith/fast_lin_arith.ML
author wenzelm
Wed, 26 Nov 2014 20:05:34 +0100
changeset 59058 a78612c67ec0
parent 58839 ccda99401bc8
child 59498 50b60f501b05
permissions -rw-r--r--
renamed "pairself" to "apply2", in accordance to @{apply 2};
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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: 31102
diff changeset
     2
    Author:     Tobias Nipkow and Tjark Weber and Sascha Boehme
6102
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
     3
46709
65a9b30bff00 clarified prems_lin_arith_tac, with subtle change of semantics: structured prems are inserted as well;
wenzelm
parents: 44654
diff changeset
     4
A generic linear arithmetic package.
6102
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
     5
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
     6
Only take premises and conclusions into account that are already
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
     7
(negated) (in)equations. lin_arith_simproc tries to prove or disprove
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
     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
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    13
signature LIN_ARITH_LOGIC =
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    14
sig
20276
d94dc40673b1 possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
webertj
parents: 20268
diff 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: 20268
diff 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: 20268
diff 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: 20268
diff 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: 20268
diff 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: 20268
diff changeset
    20
  val sym         : thm  (* x = y ==> y = x *)
31510
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff 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: 20268
diff 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: 20268
diff 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: 20268
diff 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: 20268
diff 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: 20268
diff 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: 20268
diff 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: 20268
diff changeset
    28
  val mk_nat_thm  : theory -> term -> thm
6102
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    29
end;
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    30
(*
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    31
mk_Eq(~in) = `in == False'
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    32
mk_Eq(in) = `in == True'
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    33
where `in' is an (in)equality.
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    34
23190
d45c4d6c5f15 fixed handling of meta-logic propositions
webertj
parents: 23063
diff changeset
    35
neg_prop(t) = neg  if t is wrapped up in Trueprop and neg is the
d45c4d6c5f15 fixed handling of meta-logic propositions
webertj
parents: 23063
diff changeset
    36
  (logically) negated version of t (again wrapped up in Trueprop),
d45c4d6c5f15 fixed handling of meta-logic propositions
webertj
parents: 23063
diff changeset
    37
  where the negation of a negative term is the term itself (no
d45c4d6c5f15 fixed handling of meta-logic propositions
webertj
parents: 23063
diff changeset
    38
  double negation!); raises TERM ("neg_prop", [t]) if t is not of
d45c4d6c5f15 fixed handling of meta-logic propositions
webertj
parents: 23063
diff changeset
    39
  the form 'Trueprop $ _'
6128
2acc5d36610c More arith refinements.
nipkow
parents: 6110
diff changeset
    40
2acc5d36610c More arith refinements.
nipkow
parents: 6110
diff changeset
    41
is_nat(parameter-types,t) =  t:nat
2acc5d36610c More arith refinements.
nipkow
parents: 6110
diff changeset
    42
mk_nat_thm(t) = "0 <= t"
6102
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    43
*)
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    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
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
    47
  (*internal representation of linear (in-)equations:*)
26945
9cd13e810998 renamed type decompT to decomp;
wenzelm
parents: 26835
diff changeset
    48
  type decomp = (term * Rat.rat) list * Rat.rat * string * (term * Rat.rat) list * Rat.rat * bool
9cd13e810998 renamed type decompT to decomp;
wenzelm
parents: 26835
diff changeset
    49
  val decomp: Proof.context -> term -> decomp option
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
    50
  val domain_is_nat: term -> bool
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
    51
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
    52
  (*preprocessing, performed on a representation of subgoals as list of premises:*)
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
    53
  val pre_decomp: Proof.context -> typ list * term list -> (typ list * term list) list
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
    54
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
    55
  (*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: 49387
diff changeset
    56
  val pre_tac: Proof.context -> int -> tactic
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
    57
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
    58
  (*the limit on the number of ~= allowed; because each ~= is split
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
    59
    into two cases, this can lead to an explosion*)
44654
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
    60
  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: 43568
diff changeset
    61
44654
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
    62
  val verbose: bool Config.T
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
    63
  val trace: bool Config.T
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    64
end;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    65
(*
7551
8e934d1a9ac6 Now distinguishes discrete from non-distrete types.
nipkow
parents: 6128
diff changeset
    66
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
    67
   where Rel is one of "<", "~<", "<=", "~<=" and "=" and
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
    68
         p (q, respectively) is the decomposition of the sum term x
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
    69
         (y, respectively) into a list of summand * multiplicity
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
    70
         pairs and a constant summand and d indicates if the domain
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
    71
         is discrete.
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
    72
20276
d94dc40673b1 possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
webertj
parents: 20268
diff changeset
    73
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: 20268
diff changeset
    74
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
    75
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: 19618
diff changeset
    76
internal representation of a subgoal and the corresponding theorem must
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
    77
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: 19618
diff changeset
    78
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: 19618
diff changeset
    79
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: 19618
diff changeset
    80
insensitive to them.)
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    81
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49387
diff changeset
    82
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
    83
   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
    84
   otherwise <= can grow to massive proportions.
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    85
*)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    86
6062
ede9fea461f5 Small mods.
nipkow
parents: 6056
diff changeset
    87
signature FAST_LIN_ARITH =
ede9fea461f5 Small mods.
nipkow
parents: 6056
diff changeset
    88
sig
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49387
diff changeset
    89
  val prems_lin_arith_tac: Proof.context -> int -> tactic
31102
f1e3100e6c49 mk_number replaces number_of
haftmann
parents: 30687
diff changeset
    90
  val lin_arith_tac: Proof.context -> bool -> int -> tactic
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49387
diff changeset
    91
  val lin_arith_simproc: Proof.context -> term -> thm 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: 38762
diff changeset
    92
  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: 38762
diff changeset
    93
    ({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: 49387
diff changeset
    94
      lessD: thm list, neqE: thm list, simpset: simpset,
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: 38762
diff changeset
    95
      number_of: (theory -> typ -> int -> cterm) option} ->
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: 38762
diff changeset
    96
     {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: 49387
diff changeset
    97
      lessD: thm list, neqE: thm list, simpset: simpset,
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: 38762
diff changeset
    98
      number_of: (theory -> typ -> int -> cterm) option}) ->
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: 38762
diff changeset
    99
      Context.generic -> Context.generic
38762
996afaa9254a slightly more abstract data handling in Fast_Lin_Arith;
wenzelm
parents: 38052
diff changeset
   100
  val add_inj_thms: thm list -> Context.generic -> Context.generic
996afaa9254a slightly more abstract data handling in Fast_Lin_Arith;
wenzelm
parents: 38052
diff changeset
   101
  val add_lessD: thm -> Context.generic -> Context.generic
996afaa9254a slightly more abstract data handling in Fast_Lin_Arith;
wenzelm
parents: 38052
diff changeset
   102
  val add_simps: thm list -> Context.generic -> Context.generic
996afaa9254a slightly more abstract data handling in Fast_Lin_Arith;
wenzelm
parents: 38052
diff changeset
   103
  val add_simprocs: simproc list -> Context.generic -> Context.generic
996afaa9254a slightly more abstract data handling in Fast_Lin_Arith;
wenzelm
parents: 38052
diff changeset
   104
  val set_number_of: (theory -> typ -> int -> cterm) -> Context.generic -> Context.generic
6062
ede9fea461f5 Small mods.
nipkow
parents: 6056
diff changeset
   105
end;
ede9fea461f5 Small mods.
nipkow
parents: 6056
diff changeset
   106
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   107
functor Fast_Lin_Arith
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   108
  (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
   109
struct
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   110
9420
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   111
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   112
(** theory data **)
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   113
33519
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 33317
diff changeset
   114
structure Data = Generic_Data
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22596
diff changeset
   115
(
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   116
  type T =
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   117
   {add_mono_thms: thm list,
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   118
    mult_mono_thms: thm list,
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   119
    inj_thms: thm list,
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   120
    lessD: thm list,
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   121
    neqE: thm list,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49387
diff changeset
   122
    simpset: simpset,
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: 38762
diff changeset
   123
    number_of: (theory -> typ -> int -> cterm) option};
9420
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   124
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: 38762
diff changeset
   125
  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: 38762
diff changeset
   126
   {add_mono_thms = [], mult_mono_thms = [], inj_thms = [],
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49387
diff changeset
   127
    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: 38762
diff changeset
   128
    number_of = NONE};
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16358
diff changeset
   129
  val extend = I;
33519
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 33317
diff changeset
   130
  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: 38762
diff changeset
   131
    ({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: 38762
diff changeset
   132
      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: 38762
diff changeset
   133
     {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: 38762
diff changeset
   134
      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: 23577
diff changeset
   135
    {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: 23577
diff changeset
   136
     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: 23577
diff changeset
   137
     inj_thms = Thm.merge_thms (inj_thms1, inj_thms2),
273698405054 renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents: 23577
diff changeset
   138
     lessD = Thm.merge_thms (lessD1, lessD2),
273698405054 renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents: 23577
diff changeset
   139
     neqE = Thm.merge_thms (neqE1, neqE2),
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49387
diff changeset
   140
     simpset = merge_ss (simpset1, simpset2),
41493
f05976d69141 added merge_options convenience;
wenzelm
parents: 40316
diff changeset
   141
     number_of = merge_options (number_of1, number_of2)};
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22596
diff changeset
   142
);
9420
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   143
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   144
val map_data = Data.map;
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   145
val get_data = Data.get o Context.Proof;
9420
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   146
38762
996afaa9254a slightly more abstract data handling in Fast_Lin_Arith;
wenzelm
parents: 38052
diff changeset
   147
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: 38052
diff changeset
   148
  {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: 38052
diff changeset
   149
    lessD = lessD, neqE = neqE, simpset = simpset, number_of = number_of};
996afaa9254a slightly more abstract data handling in Fast_Lin_Arith;
wenzelm
parents: 38052
diff changeset
   150
996afaa9254a slightly more abstract data handling in Fast_Lin_Arith;
wenzelm
parents: 38052
diff changeset
   151
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: 38052
diff changeset
   152
  {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: 38052
diff changeset
   153
    lessD = f lessD, neqE = neqE, simpset = simpset, number_of = number_of};
996afaa9254a slightly more abstract data handling in Fast_Lin_Arith;
wenzelm
parents: 38052
diff changeset
   154
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49387
diff changeset
   155
fun map_simpset f context =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49387
diff changeset
   156
  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: 49387
diff changeset
   157
    {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: 49387
diff changeset
   158
      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: 49387
diff changeset
   159
      number_of = number_of}) context;
38762
996afaa9254a slightly more abstract data handling in Fast_Lin_Arith;
wenzelm
parents: 38052
diff changeset
   160
996afaa9254a slightly more abstract data handling in Fast_Lin_Arith;
wenzelm
parents: 38052
diff changeset
   161
fun add_inj_thms thms = map_data (map_inj_thms (append thms));
996afaa9254a slightly more abstract data handling in Fast_Lin_Arith;
wenzelm
parents: 38052
diff changeset
   162
fun add_lessD thm = map_data (map_lessD (fn thms => thms @ [thm]));
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49387
diff changeset
   163
fun add_simps thms = map_simpset (fn ctxt => ctxt addsimps thms);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49387
diff changeset
   164
fun add_simprocs procs = map_simpset (fn ctxt => ctxt addsimprocs procs);
38762
996afaa9254a slightly more abstract data handling in Fast_Lin_Arith;
wenzelm
parents: 38052
diff changeset
   165
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: 38762
diff changeset
   166
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: 38762
diff changeset
   167
  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: 38762
diff changeset
   168
   {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: 38762
diff changeset
   169
    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: 38762
diff changeset
   170
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: 38762
diff changeset
   171
fun number_of ctxt =
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: 38762
diff changeset
   172
  (case Data.get (Context.Proof ctxt) of
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 41493
diff changeset
   173
    {number_of = SOME f, ...} => f (Proof_Context.theory_of 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: 38762
diff changeset
   174
  | _ => 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: 38762
diff changeset
   175
9420
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   176
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   177
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   178
(*** A fast decision procedure ***)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   179
(*** Code ported from HOL Light ***)
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   180
(* possible optimizations:
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   181
   use (var,coeff) rep or vector rep  tp save space;
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   182
   treat non-negative atoms separately rather than adding 0 <= atom
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   183
*)
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   184
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   185
datatype lineq_type = Eq | Le | Lt;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   186
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   187
datatype injust = Asm of int
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   188
                | Nat of int (* index of atom *)
6128
2acc5d36610c More arith refinements.
nipkow
parents: 6110
diff changeset
   189
                | LessD of injust
2acc5d36610c More arith refinements.
nipkow
parents: 6110
diff changeset
   190
                | NotLessD of injust
2acc5d36610c More arith refinements.
nipkow
parents: 6110
diff changeset
   191
                | NotLeD of injust
7551
8e934d1a9ac6 Now distinguishes discrete from non-distrete types.
nipkow
parents: 6128
diff changeset
   192
                | NotLeDD of injust
24630
351a308ab58d simplified type int (eliminated IntInf.int, integer);
wenzelm
parents: 24112
diff changeset
   193
                | Multiplied of int * injust
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   194
                | Added of injust * injust;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   195
24630
351a308ab58d simplified type int (eliminated IntInf.int, integer);
wenzelm
parents: 24112
diff changeset
   196
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
   197
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   198
(* ------------------------------------------------------------------------- *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   199
(* Finding a (counter) example from the trace of a failed elimination        *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   200
(* ------------------------------------------------------------------------- *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   201
(* Examples are represented as rational numbers,                             *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   202
(* Dont blame John Harrison for this code - it is entirely mine. TN          *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   203
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   204
exception NoEx;
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   205
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   206
(* 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: 14360
diff changeset
   207
   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: 14360
diff changeset
   208
   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: 14360
diff changeset
   209
*)
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   210
22950
haftmann
parents: 22846
diff changeset
   211
fun elim_eqns (Lineq (i, Le, cs, _)) = [(i, true, cs)]
haftmann
parents: 22846
diff changeset
   212
  | elim_eqns (Lineq (i, Eq, cs, _)) = [(i, true, cs),(~i, true, map ~ cs)]
haftmann
parents: 22846
diff changeset
   213
  | elim_eqns (Lineq (i, Lt, cs, _)) = [(i, false, cs)];
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   214
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   215
(* PRE: ex[v] must be 0! *)
24630
351a308ab58d simplified type int (eliminated IntInf.int, integer);
wenzelm
parents: 24112
diff changeset
   216
fun eval ex v (a, le, cs) =
22950
haftmann
parents: 22846
diff changeset
   217
  let
haftmann
parents: 22846
diff changeset
   218
    val rs = map Rat.rat_of_int cs;
haftmann
parents: 22846
diff changeset
   219
    val rsum = fold2 (Rat.add oo Rat.mult) rs ex Rat.zero;
23063
haftmann
parents: 23025
diff changeset
   220
  in (Rat.mult (Rat.add (Rat.rat_of_int a) (Rat.neg rsum)) (Rat.inv (nth rs v)), le) end;
haftmann
parents: 23025
diff changeset
   221
(* If nth rs v < 0, le should be negated.
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   222
   Instead this swap is taken into account in ratrelmin2.
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   223
*)
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   224
22950
haftmann
parents: 22846
diff changeset
   225
fun ratrelmin2 (x as (r, ler), y as (s, les)) =
23520
483fe92f00c1 tuned arithmetic modules
haftmann
parents: 23297
diff changeset
   226
  case Rat.ord (r, s)
22950
haftmann
parents: 22846
diff changeset
   227
   of EQUAL => (r, (not ler) andalso (not les))
haftmann
parents: 22846
diff changeset
   228
    | LESS => x
haftmann
parents: 22846
diff changeset
   229
    | GREATER => y;
haftmann
parents: 22846
diff changeset
   230
haftmann
parents: 22846
diff changeset
   231
fun ratrelmax2 (x as (r, ler), y as (s, les)) =
23520
483fe92f00c1 tuned arithmetic modules
haftmann
parents: 23297
diff changeset
   232
  case Rat.ord (r, s)
22950
haftmann
parents: 22846
diff changeset
   233
   of EQUAL => (r, ler andalso les)
haftmann
parents: 22846
diff changeset
   234
    | LESS => y
haftmann
parents: 22846
diff changeset
   235
    | GREATER => x;
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   236
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   237
val ratrelmin = foldr1 ratrelmin2;
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   238
val ratrelmax = foldr1 ratrelmax2;
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   239
22950
haftmann
parents: 22846
diff changeset
   240
fun ratexact up (r, exact) =
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   241
  if exact then r else
22950
haftmann
parents: 22846
diff changeset
   242
  let
38052
04a8de29e8f7 dropped dead code
haftmann
parents: 36945
diff changeset
   243
    val (_, q) = Rat.quotient_of_rat r;
22950
haftmann
parents: 22846
diff changeset
   244
    val nth = Rat.inv (Rat.rat_of_int q);
haftmann
parents: 22846
diff changeset
   245
  in Rat.add r (if up then nth else Rat.neg nth) end;
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   246
22950
haftmann
parents: 22846
diff changeset
   247
fun ratmiddle (r, s) = Rat.mult (Rat.add r s) (Rat.inv Rat.two);
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   248
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   249
fun choose2 d ((lb, exactl), (ub, exactu)) =
23520
483fe92f00c1 tuned arithmetic modules
haftmann
parents: 23297
diff changeset
   250
  let val ord = Rat.sign lb in
22950
haftmann
parents: 22846
diff changeset
   251
  if (ord = LESS orelse exactl) andalso (ord = GREATER orelse exactu)
haftmann
parents: 22846
diff changeset
   252
    then Rat.zero
haftmann
parents: 22846
diff changeset
   253
    else if not d then
haftmann
parents: 22846
diff changeset
   254
      if ord = GREATER
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   255
        then if exactl then lb else ratmiddle (lb, ub)
22950
haftmann
parents: 22846
diff changeset
   256
        else if exactu then ub else ratmiddle (lb, ub)
haftmann
parents: 22846
diff changeset
   257
      else (*discrete domain, both bounds must be exact*)
23025
7507f94adc32 dropped nonsense comment
haftmann
parents: 22950
diff changeset
   258
      if ord = GREATER
22950
haftmann
parents: 22846
diff changeset
   259
        then let val lb' = Rat.roundup lb in
haftmann
parents: 22846
diff changeset
   260
          if Rat.le lb' ub then lb' else raise NoEx end
haftmann
parents: 22846
diff changeset
   261
        else let val ub' = Rat.rounddown ub in
haftmann
parents: 22846
diff changeset
   262
          if Rat.le lb ub' then ub' else raise NoEx end
haftmann
parents: 22846
diff changeset
   263
  end;
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   264
22950
haftmann
parents: 22846
diff changeset
   265
fun findex1 discr (v, lineqs) ex =
haftmann
parents: 22846
diff changeset
   266
  let
23063
haftmann
parents: 23025
diff changeset
   267
    val nz = filter (fn (Lineq (_, _, cs, _)) => nth cs v <> 0) lineqs;
22950
haftmann
parents: 22846
diff changeset
   268
    val ineqs = maps elim_eqns nz;
23063
haftmann
parents: 23025
diff changeset
   269
    val (ge, le) = List.partition (fn (_,_,cs) => nth cs v > 0) ineqs
22950
haftmann
parents: 22846
diff changeset
   270
    val lb = ratrelmax (map (eval ex v) ge)
haftmann
parents: 22846
diff changeset
   271
    val ub = ratrelmin (map (eval ex v) le)
21109
f8f89be75e81 dropped nth_update
haftmann
parents: 20433
diff changeset
   272
  in nth_map v (K (choose2 (nth discr v) (lb, ub))) ex end;
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   273
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   274
fun elim1 v x =
23063
haftmann
parents: 23025
diff changeset
   275
  map (fn (a,le,bs) => (Rat.add a (Rat.neg (Rat.mult (nth bs v) x)), le,
21109
f8f89be75e81 dropped nth_update
haftmann
parents: 20433
diff changeset
   276
                        nth_map v (K Rat.zero) bs));
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   277
23520
483fe92f00c1 tuned arithmetic modules
haftmann
parents: 23297
diff changeset
   278
fun single_var v (_, _, cs) = case filter_out (curry (op =) EQUAL o Rat.sign) cs
23063
haftmann
parents: 23025
diff changeset
   279
 of [x] => x =/ nth cs v
haftmann
parents: 23025
diff changeset
   280
  | _ => false;
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   281
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   282
(* The base case:
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   283
   all variables occur only with positive or only with negative coefficients *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   284
fun pick_vars discr (ineqs,ex) =
23520
483fe92f00c1 tuned arithmetic modules
haftmann
parents: 23297
diff changeset
   285
  let val nz = filter_out (fn (_,_,cs) => forall (curry (op =) EQUAL o Rat.sign) cs) ineqs
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   286
  in case nz of [] => ex
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   287
     | (_,_,cs) :: _ =>
23520
483fe92f00c1 tuned arithmetic modules
haftmann
parents: 23297
diff changeset
   288
       let val v = find_index (not o curry (op =) EQUAL o Rat.sign) cs
22950
haftmann
parents: 22846
diff changeset
   289
           val d = nth discr v;
23520
483fe92f00c1 tuned arithmetic modules
haftmann
parents: 23297
diff changeset
   290
           val pos = not (Rat.sign (nth cs v) = LESS);
22950
haftmann
parents: 22846
diff changeset
   291
           val sv = filter (single_var v) nz;
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   292
           val minmax =
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   293
             if pos then if d then Rat.roundup o fst o ratrelmax
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   294
                         else ratexact true o ratrelmax
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   295
                    else if d then Rat.rounddown o fst o ratrelmin
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   296
                         else ratexact false o ratrelmin
23063
haftmann
parents: 23025
diff changeset
   297
           val bnds = map (fn (a,le,bs) => (Rat.mult a (Rat.inv (nth bs v)), le)) sv
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   298
           val x = minmax((Rat.zero,if pos then true else false)::bnds)
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   299
           val ineqs' = elim1 v x nz
21109
f8f89be75e81 dropped nth_update
haftmann
parents: 20433
diff changeset
   300
           val ex' = nth_map v (K x) ex
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   301
       in pick_vars discr (ineqs',ex') end
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   302
  end;
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   303
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   304
fun findex0 discr n lineqs =
22950
haftmann
parents: 22846
diff changeset
   305
  let val ineqs = maps elim_eqns lineqs
haftmann
parents: 22846
diff changeset
   306
      val rineqs = map (fn (a,le,cs) => (Rat.rat_of_int a, le, map Rat.rat_of_int cs))
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   307
                       ineqs
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   308
  in pick_vars discr (rineqs,replicate n Rat.zero) end;
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   309
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   310
(* ------------------------------------------------------------------------- *)
23197
f96d909eda37 cosmetic
webertj
parents: 23195
diff changeset
   311
(* End of counterexample finder. The actual decision procedure starts here.  *)
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   312
(* ------------------------------------------------------------------------- *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   313
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   314
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   315
(* 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
   316
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   317
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   318
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
   319
  | find_add_type(x,Eq) = x
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   320
  | find_add_type(_,Lt) = Lt
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   321
  | find_add_type(Lt,_) = Lt
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   322
  | find_add_type(Le,Le) = Le;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   323
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   324
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   325
(* Multiply out an (in)equation.                                             *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   326
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   327
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   328
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
   329
  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: 38763
diff changeset
   330
  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: 38763
diff changeset
   331
  else if n < 0 andalso (ty=Le orelse ty=Lt) then raise Fail "multiply_ineq"
33002
f3f02f36a3e2 uniform use of Integer.add/mult/sum/prod;
wenzelm
parents: 32952
diff changeset
   332
  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
   333
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   334
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   335
(* Add together (in)equations.                                               *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   336
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   337
38052
04a8de29e8f7 dropped dead code
haftmann
parents: 36945
diff changeset
   338
fun add_ineq (Lineq (k1,ty1,l1,just1)) (Lineq (k2,ty2,l2,just2)) =
33002
f3f02f36a3e2 uniform use of Integer.add/mult/sum/prod;
wenzelm
parents: 32952
diff changeset
   339
  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
   340
  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
   341
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   342
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   343
(* 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
   344
(* 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
   345
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   346
49387
167708456269 dropped some unused identifiers
haftmann
parents: 47060
diff changeset
   347
fun elim_var v (i1 as Lineq(_,ty1,l1,_)) (i2 as Lineq(_,ty2,l2,_)) =
23063
haftmann
parents: 23025
diff changeset
   348
  let val c1 = nth l1 v and c2 = nth l2 v
23261
85f27f79232f tuned integers
haftmann
parents: 23197
diff changeset
   349
      val m = Integer.lcm (abs c1) (abs c2)
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   350
      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
   351
      val (n1,n2) =
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   352
        if (c1 >= 0) = (c2 >= 0)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   353
        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
   354
             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: 38763
diff changeset
   355
                  else raise Fail "elim_var"
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   356
        else (m1,m2)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   357
      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
   358
                    then (~n1,~n2) else (n1,n2)
31510
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   359
  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
   360
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   361
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   362
(* The main refutation-finding code.                                         *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   363
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   364
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   365
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
   366
38052
04a8de29e8f7 dropped dead code
haftmann
parents: 36945
diff changeset
   367
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
   368
  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
   369
24630
351a308ab58d simplified type int (eliminated IntInf.int, integer);
wenzelm
parents: 24112
diff changeset
   370
fun calc_blowup l =
33317
b4534348b8fd standardized filter/filter_out;
wenzelm
parents: 33245
diff changeset
   371
  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: 24112
diff changeset
   372
  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
   373
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   374
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   375
(* Main elimination code:                                                    *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   376
(*                                                                           *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   377
(* (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
   378
(*                                                                           *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   379
(* (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
   380
(* 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
   381
(*                                                                           *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   382
(* (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
   383
(* 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
   384
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   385
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   386
fun extract_first p =
31510
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   387
  let
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   388
    fun extract xs (y::ys) = if p y then (y, xs @ ys) else extract (y::xs) ys
47060
e2741ec9ae36 prefer explicitly qualified exception List.Empty;
wenzelm
parents: 46709
diff changeset
   389
      | extract xs [] = raise List.Empty
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   390
  in extract [] end;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   391
44654
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   392
fun print_ineqs ctxt ineqs =
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   393
  if Config.get ctxt LA_Data.trace then
12262
11ff5f47df6e use tracing function for trace output;
wenzelm
parents: 12109
diff changeset
   394
     tracing(cat_lines(""::map (fn Lineq(c,t,l,_) =>
24630
351a308ab58d simplified type int (eliminated IntInf.int, integer);
wenzelm
parents: 24112
diff changeset
   395
       string_of_int c ^
9073
40d8dfac96b8 tracing flag for arith_tac
paulson
parents: 8263
diff changeset
   396
       (case t of Eq => " =  " | Lt=> " <  " | Le => " <= ") ^
24630
351a308ab58d simplified type int (eliminated IntInf.int, integer);
wenzelm
parents: 24112
diff changeset
   397
       commas(map string_of_int l)) ineqs))
9073
40d8dfac96b8 tracing flag for arith_tac
paulson
parents: 8263
diff changeset
   398
  else ();
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   399
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   400
type history = (int * lineq list) list;
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   401
datatype result = Success of injust | Failure of history;
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   402
44654
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   403
fun elim ctxt (ineqs, hist) =
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   404
  let val _ = print_ineqs ctxt ineqs
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   405
      val (triv, nontriv) = List.partition is_trivial ineqs in
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   406
  if not (null triv)
31510
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   407
  then case Library.find_first is_contradictory triv of
44654
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   408
         NONE => elim ctxt (nontriv, hist)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15184
diff changeset
   409
       | SOME(Lineq(_,_,_,j)) => Success j
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   410
  else
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   411
  if null nontriv then Failure hist
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   412
  else
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   413
  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: 19618
diff changeset
   414
  if not (null eqs) then
31510
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   415
     let val c =
33042
ddf1f03a9ad9 curried union as canonical list operation
haftmann
parents: 33038
diff changeset
   416
           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: 31102
diff changeset
   417
           |> filter (fn i => i <> 0)
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58839
diff changeset
   418
           |> sort (int_ord o apply2 abs)
31510
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   419
           |> hd
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   420
         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: 35872
diff changeset
   421
               extract_first (fn Lineq(_,_,l,_) => member (op =) l c) eqs
31986
a68f88d264f7 dropped find_index_eq
haftmann
parents: 31638
diff changeset
   422
         val v = find_index (fn v => v = c) ceq
23063
haftmann
parents: 23025
diff changeset
   423
         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
   424
                                     (othereqs @ noneqs)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   425
         val others = map (elim_var v eq) roth @ ioth
44654
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   426
     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
   427
  else
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   428
  let val lists = map (fn (Lineq(_,_,l,_)) => l) noneqs
23063
haftmann
parents: 23025
diff changeset
   429
      val numlist = 0 upto (length (hd lists) - 1)
haftmann
parents: 23025
diff changeset
   430
      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
   431
      val blows = map calc_blowup coeffs
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   432
      val iblows = blows ~~ numlist
23063
haftmann
parents: 23025
diff changeset
   433
      val nziblows = filter_out (fn (i, _) => i = 0) iblows
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   434
  in if null nziblows then Failure((~1,nontriv)::hist)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   435
     else
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   436
     let val (c,v) = hd(sort (fn (x,y) => int_ord(fst(x),fst(y))) nziblows)
23063
haftmann
parents: 23025
diff changeset
   437
         val (no,yes) = List.partition (fn (Lineq(_,_,l,_)) => nth l v = 0) ineqs
haftmann
parents: 23025
diff changeset
   438
         val (pos,neg) = List.partition(fn (Lineq(_,_,l,_)) => nth l v > 0) yes
44654
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   439
     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
   440
  end
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   441
  end
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   442
  end;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   443
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   444
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   445
(* Translate back a proof.                                                   *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   446
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   447
44654
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   448
fun trace_thm ctxt msgs th =
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   449
 (if Config.get ctxt LA_Data.trace
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   450
  then tracing (cat_lines (msgs @ [Display.string_of_thm ctxt th]))
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   451
  else (); th);
9073
40d8dfac96b8 tracing flag for arith_tac
paulson
parents: 8263
diff changeset
   452
44654
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   453
fun trace_term ctxt msgs t =
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   454
 (if Config.get ctxt LA_Data.trace
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   455
  then tracing (cat_lines (msgs @ [Syntax.string_of_term ctxt t]))
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   456
  else (); t);
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   457
44654
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   458
fun trace_msg ctxt msg =
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   459
  if Config.get ctxt LA_Data.trace then tracing msg else ();
9073
40d8dfac96b8 tracing flag for arith_tac
paulson
parents: 8263
diff changeset
   460
52131
366fa32ee2a3 tuned signature;
wenzelm
parents: 51930
diff changeset
   461
val union_term = union Envir.aeconv;
366fa32ee2a3 tuned signature;
wenzelm
parents: 51930
diff changeset
   462
val union_bterm = union (fn ((b:bool, t), (b', t')) => b = b' andalso Envir.aeconv (t, t'));
26835
404550067389 Lookup and union operations on terms are now modulo eta conversion.
berghofe
parents: 24920
diff changeset
   463
31510
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   464
fun add_atoms (lhs, _, _, rhs, _, _) =
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   465
  union_term (map fst lhs) o union_term (map fst rhs);
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   466
31510
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   467
fun atoms_of ds = fold add_atoms ds [];
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   468
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   469
(*
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   470
Simplification may detect a contradiction 'prematurely' due to type
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   471
information: n+1 <= 0 is simplified to False and does not need to be crossed
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   472
with 0 <= n.
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   473
*)
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   474
local
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   475
  exception FalseE of thm
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   476
in
27020
b5b8afc9fdcd added warning_count for issued reconstruction failure messages (limit 10);
wenzelm
parents: 26945
diff changeset
   477
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49387
diff changeset
   478
fun mkthm ctxt asms (just: injust) =
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   479
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 41493
diff changeset
   480
    val thy = Proof_Context.theory_of 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: 38762
diff changeset
   481
    val {add_mono_thms, mult_mono_thms, inj_thms, lessD, simpset, ...} = get_data ctxt;
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: 38762
diff changeset
   482
    val number_of = number_of ctxt;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49387
diff changeset
   483
    val simpset_ctxt = put_simpset simpset ctxt;
31510
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   484
    fun only_concl f thm =
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   485
      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: 31102
diff changeset
   486
    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: 31102
diff changeset
   487
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   488
    fun use_first rules thm =
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   489
      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: 31102
diff changeset
   490
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   491
    fun add2 thm1 thm2 =
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   492
      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: 31102
diff changeset
   493
    fun try_add thms thm = get_first (fn th => add2 th thm) thms;
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   494
31510
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   495
    fun add_thms thm1 thm2 =
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   496
      (case add2 thm1 thm2 of
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   497
        NONE =>
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   498
          (case try_add ([thm1] RL inj_thms) thm2 of
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   499
            NONE =>
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   500
              (the (try_add ([thm2] RL inj_thms) thm1)
51930
52fd62618631 prefer explicitly qualified exceptions, which is particular important for robust handlers;
wenzelm
parents: 51717
diff changeset
   501
                handle Option.Option =>
44654
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   502
                  (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: 38763
diff changeset
   503
                   raise Fail "Linear arithmetic: failed to add thms"))
31510
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   504
          | SOME thm => thm)
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   505
      | SOME thm => thm);
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   506
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   507
    fun mult_by_add n thm =
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   508
      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: 31102
diff changeset
   509
      in mul n thm end;
10575
c78d26d5c3c1 Now adjusted to mixed terms involving coercions.
nipkow
parents: 9420
diff changeset
   510
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49387
diff changeset
   511
    val rewr = Simplifier.rewrite simpset_ctxt;
31510
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   512
    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: 31102
diff changeset
   513
      (Conv.binop_conv rewr)));
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   514
    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: 31102
diff changeset
   515
      let val cv = Conv.arg1_conv (Conv.arg_conv rewr)
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   516
      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: 19618
diff changeset
   517
31510
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   518
    fun mult n thm =
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   519
      (case use_first mult_mono_thms thm of
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   520
        NONE => mult_by_add n thm
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   521
      | SOME mth =>
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   522
          let
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   523
            val cv = mth |> Thm.cprop_of |> Drule.strip_imp_concl
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   524
              |> Thm.dest_arg |> Thm.dest_arg1 |> Thm.dest_arg1
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   525
            val T = #T (Thm.rep_cterm cv)
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   526
          in
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   527
            mth
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: 38762
diff changeset
   528
            |> Thm.instantiate ([], [(cv, number_of T n)])
31510
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   529
            |> rewrite_concl
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   530
            |> discharge_prem
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   531
            handle CTERM _ => mult_by_add n thm
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   532
                 | THM _ => mult_by_add n thm
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   533
          end);
10691
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
   534
31510
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   535
    fun mult_thm (n, thm) =
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   536
      if n = ~1 then thm RS LA_Logic.sym
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   537
      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: 31102
diff changeset
   538
      else mult n thm;
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   539
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   540
    fun simp thm =
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49387
diff changeset
   541
      let val thm' = trace_thm ctxt ["Simplified:"] (full_simplify simpset_ctxt thm)
31510
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   542
      in if LA_Logic.is_False thm' then raise FalseE thm' else thm' end;
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   543
44654
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   544
    fun mk (Asm i) = trace_thm ctxt ["Asm " ^ string_of_int i] (nth asms i)
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   545
      | mk (Nat i) = trace_thm ctxt ["Nat " ^ string_of_int i] (LA_Logic.mk_nat_thm thy (nth atoms i))
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   546
      | mk (LessD j) = trace_thm ctxt ["L"] (hd ([mk j] RL lessD))
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   547
      | mk (NotLeD j) = trace_thm ctxt ["NLe"] (mk j RS LA_Logic.not_leD)
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   548
      | mk (NotLeDD j) = trace_thm ctxt ["NLeD"] (hd ([mk j RS LA_Logic.not_leD] RL lessD))
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   549
      | mk (NotLessD j) = trace_thm ctxt ["NL"] (mk j RS LA_Logic.not_lessD)
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   550
      | mk (Added (j1, j2)) = simp (trace_thm ctxt ["+"] (add_thms (mk j1) (mk j2)))
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: 31986
diff changeset
   551
      | mk (Multiplied (n, j)) =
44654
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   552
          (trace_msg ctxt ("*" ^ string_of_int n); trace_thm ctxt ["*"] (mult_thm (n, mk j)))
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   553
27020
b5b8afc9fdcd added warning_count for issued reconstruction failure messages (limit 10);
wenzelm
parents: 26945
diff changeset
   554
  in
b5b8afc9fdcd added warning_count for issued reconstruction failure messages (limit 10);
wenzelm
parents: 26945
diff changeset
   555
    let
44654
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   556
      val _ = trace_msg ctxt "mkthm";
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   557
      val thm = trace_thm ctxt ["Final thm:"] (mk just);
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49387
diff changeset
   558
      val fls = simplify simpset_ctxt thm;
44654
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   559
      val _ = trace_thm ctxt ["After simplification:"] fls;
27020
b5b8afc9fdcd added warning_count for issued reconstruction failure messages (limit 10);
wenzelm
parents: 26945
diff changeset
   560
      val _ =
b5b8afc9fdcd added warning_count for issued reconstruction failure messages (limit 10);
wenzelm
parents: 26945
diff changeset
   561
        if LA_Logic.is_False fls then ()
b5b8afc9fdcd added warning_count for issued reconstruction failure messages (limit 10);
wenzelm
parents: 26945
diff changeset
   562
        else
35872
9b579860d59b removed warning_count (known causes for warnings have been resolved)
boehmes
parents: 35861
diff changeset
   563
         (tracing (cat_lines
9b579860d59b removed warning_count (known causes for warnings have been resolved)
boehmes
parents: 35861
diff changeset
   564
           (["Assumptions:"] @ map (Display.string_of_thm ctxt) asms @ [""] @
9b579860d59b removed warning_count (known causes for warnings have been resolved)
boehmes
parents: 35861
diff changeset
   565
            ["Proved:", Display.string_of_thm ctxt fls, ""]));
9b579860d59b removed warning_count (known causes for warnings have been resolved)
boehmes
parents: 35861
diff changeset
   566
          warning "Linear arithmetic should have refuted the assumptions.\n\
9b579860d59b removed warning_count (known causes for warnings have been resolved)
boehmes
parents: 35861
diff changeset
   567
            \Please inform Tobias Nipkow.")
27020
b5b8afc9fdcd added warning_count for issued reconstruction failure messages (limit 10);
wenzelm
parents: 26945
diff changeset
   568
    in fls end
44654
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   569
    handle FalseE thm => trace_thm ctxt ["False reached early:"] thm
27020
b5b8afc9fdcd added warning_count for issued reconstruction failure messages (limit 10);
wenzelm
parents: 26945
diff changeset
   570
  end;
b5b8afc9fdcd added warning_count for issued reconstruction failure messages (limit 10);
wenzelm
parents: 26945
diff changeset
   571
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   572
end;
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   573
23261
85f27f79232f tuned integers
haftmann
parents: 23197
diff changeset
   574
fun coeff poly atom =
52131
366fa32ee2a3 tuned signature;
wenzelm
parents: 51930
diff changeset
   575
  AList.lookup Envir.aeconv poly atom |> the_default 0;
10691
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
   576
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
   577
fun integ(rlhs,r,rel,rrhs,s,d) =
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   578
let val (rn,rd) = Rat.quotient_of_rat r and (sn,sd) = Rat.quotient_of_rat s
24630
351a308ab58d simplified type int (eliminated IntInf.int, integer);
wenzelm
parents: 24112
diff changeset
   579
    val m = Integer.lcms(map (abs o snd o Rat.quotient_of_rat) (r :: s :: map snd rlhs @ map snd rrhs))
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22596
diff changeset
   580
    fun mult(t,r) =
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   581
        let val (i,j) = Rat.quotient_of_rat r
15965
f422f8283491 Use of IntInf.int instead of int in most numeric simprocs; avoids
paulson
parents: 15922
diff changeset
   582
        in (t,i * (m div j)) end
12932
3bda5306d262 updated debugging output
nipkow
parents: 12311
diff changeset
   583
in (m,(map mult rlhs, rn*(m div rd), rel, map mult rrhs, sn*(m div sd), d)) end
10691
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
   584
38052
04a8de29e8f7 dropped dead code
haftmann
parents: 36945
diff changeset
   585
fun mklineq atoms =
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   586
  fn (item, k) =>
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   587
  let val (m, (lhs,i,rel,rhs,j,discrete)) = integ item
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   588
      val lhsa = map (coeff lhs) atoms
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   589
      and rhsa = map (coeff rhs) atoms
18330
444f16d232a2 introduced new map2, fold
haftmann
parents: 18011
diff changeset
   590
      val diff = map2 (curry (op -)) rhsa lhsa
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   591
      val c = i-j
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   592
      val just = Asm k
31511
boehmes
parents: 31510
diff changeset
   593
      fun lineq(c,le,cs,j) = Lineq(c,le,cs, if m=1 then j else Multiplied(m,j))
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   594
  in case rel of
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   595
      "<="   => lineq(c,Le,diff,just)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   596
     | "~<=" => if discrete
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   597
                then lineq(1-c,Le,map (op ~) diff,NotLeDD(just))
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   598
                else lineq(~c,Lt,map (op ~) diff,NotLeD(just))
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   599
     | "<"   => if discrete
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   600
                then lineq(c+1,Le,diff,LessD(just))
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   601
                else lineq(c,Lt,diff,just)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   602
     | "~<"  => lineq(~c,Le,map (op~) diff,NotLessD(just))
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   603
     | "="   => 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: 38763
diff changeset
   604
     | _     => raise Fail ("mklineq" ^ rel)
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   605
  end;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   606
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   607
(* ------------------------------------------------------------------------- *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   608
(* Print (counter) example                                                   *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   609
(* ------------------------------------------------------------------------- *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   610
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   611
fun print_atom((a,d),r) =
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   612
  let val (p,q) = Rat.quotient_of_rat r
24630
351a308ab58d simplified type int (eliminated IntInf.int, integer);
wenzelm
parents: 24112
diff changeset
   613
      val s = if d then string_of_int p else
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   614
              if p = 0 then "0"
24630
351a308ab58d simplified type int (eliminated IntInf.int, integer);
wenzelm
parents: 24112
diff changeset
   615
              else string_of_int p ^ "/" ^ string_of_int q
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   616
  in a ^ " = " ^ s end;
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   617
43607
119767e1ccb4 linarith counterexamples now provide only valuations for variables (which should restrict the number of linarith trace messages);
boehmes
parents: 43568
diff changeset
   618
fun is_variable (Free _) = true
119767e1ccb4 linarith counterexamples now provide only valuations for variables (which should restrict the number of linarith trace messages);
boehmes
parents: 43568
diff changeset
   619
  | is_variable (Var _) = true
119767e1ccb4 linarith counterexamples now provide only valuations for variables (which should restrict the number of linarith trace messages);
boehmes
parents: 43568
diff changeset
   620
  | is_variable (Bound _) = true
119767e1ccb4 linarith counterexamples now provide only valuations for variables (which should restrict the number of linarith trace messages);
boehmes
parents: 43568
diff changeset
   621
  | is_variable _ = false
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   622
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   623
fun trace_ex ctxt params atoms discr n (hist: history) =
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   624
  case hist of
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   625
    [] => ()
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   626
  | (v, lineqs) :: hist' =>
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   627
      let
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   628
        val frees = map Free params
24920
2a45e400fdad generic Syntax.pretty/string_of operations;
wenzelm
parents: 24630
diff changeset
   629
        fun show_term t = Syntax.string_of_term ctxt (subst_bounds (frees, t))
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   630
        val start =
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   631
          if v = ~1 then (hist', findex0 discr n lineqs)
22950
haftmann
parents: 22846
diff changeset
   632
          else (hist, replicate n Rat.zero)
43607
119767e1ccb4 linarith counterexamples now provide only valuations for variables (which should restrict the number of linarith trace messages);
boehmes
parents: 43568
diff changeset
   633
        val produce_ex =
119767e1ccb4 linarith counterexamples now provide only valuations for variables (which should restrict the number of linarith trace messages);
boehmes
parents: 43568
diff changeset
   634
          map print_atom #> commas #>
119767e1ccb4 linarith counterexamples now provide only valuations for variables (which should restrict the number of linarith trace messages);
boehmes
parents: 43568
diff changeset
   635
          prefix "Counterexample (possibly spurious):\n"
119767e1ccb4 linarith counterexamples now provide only valuations for variables (which should restrict the number of linarith trace messages);
boehmes
parents: 43568
diff changeset
   636
        val ex = (
119767e1ccb4 linarith counterexamples now provide only valuations for variables (which should restrict the number of linarith trace messages);
boehmes
parents: 43568
diff changeset
   637
          uncurry (fold (findex1 discr)) start
119767e1ccb4 linarith counterexamples now provide only valuations for variables (which should restrict the number of linarith trace messages);
boehmes
parents: 43568
diff changeset
   638
          |> map2 pair (atoms ~~ discr)
119767e1ccb4 linarith counterexamples now provide only valuations for variables (which should restrict the number of linarith trace messages);
boehmes
parents: 43568
diff changeset
   639
          |> filter (fn ((t, _), _) => is_variable t)
119767e1ccb4 linarith counterexamples now provide only valuations for variables (which should restrict the number of linarith trace messages);
boehmes
parents: 43568
diff changeset
   640
          |> map (apfst (apfst show_term))
119767e1ccb4 linarith counterexamples now provide only valuations for variables (which should restrict the number of linarith trace messages);
boehmes
parents: 43568
diff changeset
   641
          |> (fn [] => NONE | sdss => SOME (produce_ex sdss)))
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   642
          handle NoEx => NONE
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   643
      in
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   644
        case ex of
44654
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   645
          SOME s =>
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   646
            (warning "Linear arithmetic failed -- see trace for a (potentially spurious) counterexample.";
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   647
              tracing s)
30687
16f6efd4e599 tuned error messages
haftmann
parents: 30406
diff changeset
   648
        | NONE => warning "Linear arithmetic failed"
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   649
      end;
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   650
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   651
(* ------------------------------------------------------------------------- *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   652
20268
1fe9aed8fcac code reformatted
webertj
parents: 20254
diff changeset
   653
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: 19618
diff changeset
   654
  if LA_Logic.is_nat (pTs, atom)
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   655
  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: 19618
diff changeset
   656
       in SOME (Lineq (0, Le, l, Nat i)) end
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   657
  else NONE;
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   658
13186
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   659
(* This code is tricky. It takes a list of premises in the order they occur
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15184
diff changeset
   660
in the subgoal. Numerical premises are coded as SOME(tuple), non-numerical
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15184
diff changeset
   661
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: 13105
diff changeset
   662
a Lineq. The tricky bit is to convert ~= which is split into two cases < and
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   663
>. 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: 13105
diff changeset
   664
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: 13105
diff changeset
   665
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: 13105
diff changeset
   666
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: 13105
diff changeset
   667
applies the generated refutation thms (see function 'refute_tac').
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   668
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   669
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: 13105
diff changeset
   670
*)
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   671
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   672
(* 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: 20268
diff changeset
   673
(*        could be intertwined: separate the first (fully split) case,       *)
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   674
(*        refute it, continue with splitting and refuting.  Terminate with   *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   675
(*        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: 19618
diff changeset
   676
(*        splitting until after a refutation for other cases has been found. *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   677
30406
15dc25f8a0e2 Instead of giving up entirely, arith now ignores all inequalities when there are too many.
webertj
parents: 27020
diff changeset
   678
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: 20268
diff changeset
   679
let
d94dc40673b1 possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
webertj
parents: 20268
diff changeset
   680
  (* 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: 20268
diff changeset
   681
  (* '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: 20268
diff changeset
   682
  (* level                                                          *)
d94dc40673b1 possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
webertj
parents: 20268
diff changeset
   683
  (* 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: 20268
diff changeset
   684
  (*        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: 20268
diff changeset
   685
  (*        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: 20268
diff changeset
   686
  (*        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: 20268
diff changeset
   687
  (*        can be applied, and split the premise accordingly.      *)
26945
9cd13e810998 renamed type decompT to decomp;
wenzelm
parents: 26835
diff changeset
   688
  fun elim_neq (ineqs : (LA_Data.decomp option * bool) list) :
9cd13e810998 renamed type decompT to decomp;
wenzelm
parents: 26835
diff changeset
   689
               (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: 20268
diff changeset
   690
  let
26945
9cd13e810998 renamed type decompT to decomp;
wenzelm
parents: 26835
diff changeset
   691
    fun elim_neq' nat_only ([] : (LA_Data.decomp option * bool) list) :
9cd13e810998 renamed type decompT to decomp;
wenzelm
parents: 26835
diff changeset
   692
                  (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: 20268
diff changeset
   693
          [[]]
d94dc40673b1 possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
webertj
parents: 20268
diff changeset
   694
      | 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: 20268
diff changeset
   695
          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: 20268
diff changeset
   696
      | 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: 20268
diff changeset
   697
          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: 20268
diff changeset
   698
            (* [| ?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: 20268
diff changeset
   699
            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: 20268
diff changeset
   700
            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: 20268
diff changeset
   701
          else
d94dc40673b1 possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
webertj
parents: 20268
diff changeset
   702
            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: 20268
diff changeset
   703
  in
d94dc40673b1 possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
webertj
parents: 20268
diff changeset
   704
    ineqs |> elim_neq' true
26945
9cd13e810998 renamed type decompT to decomp;
wenzelm
parents: 26835
diff changeset
   705
          |> maps (elim_neq' false)
20276
d94dc40673b1 possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
webertj
parents: 20268
diff changeset
   706
  end
13464
c98321b8d638 Fixed two bugs
nipkow
parents: 13186
diff changeset
   707
30406
15dc25f8a0e2 Instead of giving up entirely, arith now ignores all inequalities when there are too many.
webertj
parents: 27020
diff changeset
   708
  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: 27020
diff changeset
   709
    | 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: 27020
diff changeset
   710
      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: 27020
diff changeset
   711
20276
d94dc40673b1 possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
webertj
parents: 20268
diff changeset
   712
  fun number_hyps _ []             = []
d94dc40673b1 possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
webertj
parents: 20268
diff changeset
   713
    | 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: 20268
diff changeset
   714
    | 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: 20268
diff changeset
   715
d94dc40673b1 possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
webertj
parents: 20268
diff changeset
   716
  val result = (Ts, terms)
d94dc40673b1 possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
webertj
parents: 20268
diff changeset
   717
    |> (* user-defined preprocessing of the subgoal *)
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   718
       (if do_pre then LA_Data.pre_decomp ctxt else Library.single)
44654
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   719
    |> tap (fn subgoals => trace_msg ctxt ("Preprocessing yields " ^
23195
f065f7c846fe additional tracing information
webertj
parents: 23190
diff changeset
   720
         string_of_int (length subgoals) ^ " subgoal(s) total."))
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22596
diff changeset
   721
    |> (* produce the internal encoding of (in-)equalities *)
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   722
       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: 20268
diff changeset
   723
    |> (* splitting of inequalities *)
30406
15dc25f8a0e2 Instead of giving up entirely, arith now ignores all inequalities when there are too many.
webertj
parents: 27020
diff changeset
   724
       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: 27020
diff changeset
   725
                     Library.single o map ignore_neq))
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22596
diff changeset
   726
    |> 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: 20268
diff changeset
   727
    |> (* numbering of hypotheses, ignoring irrelevant ones *)
d94dc40673b1 possible disagreement between proof search and proof reconstruction when eliminating inequalities over different types fixed
webertj
parents: 20268
diff changeset
   728
       map (apsnd (number_hyps 0))
23195
f065f7c846fe additional tracing information
webertj
parents: 23190
diff changeset
   729
in
44654
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   730
  trace_msg ctxt ("Splitting of inequalities yields " ^
23195
f065f7c846fe additional tracing information
webertj
parents: 23190
diff changeset
   731
    string_of_int (length result) ^ " subgoal(s) total.");
f065f7c846fe additional tracing information
webertj
parents: 23190
diff changeset
   732
  result
f065f7c846fe additional tracing information
webertj
parents: 23190
diff changeset
   733
end;
13464
c98321b8d638 Fixed two bugs
nipkow
parents: 13186
diff changeset
   734
33245
65232054ffd0 eliminated some old folds;
wenzelm
parents: 33042
diff changeset
   735
fun add_datoms ((lhs,_,_,rhs,_,d) : LA_Data.decomp, _) (dats : (bool * term) list) =
26835
404550067389 Lookup and union operations on terms are now modulo eta conversion.
berghofe
parents: 24920
diff changeset
   736
  union_bterm (map (pair d o fst) lhs) (union_bterm (map (pair d o fst) rhs) dats);
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   737
26945
9cd13e810998 renamed type decompT to decomp;
wenzelm
parents: 26835
diff changeset
   738
fun discr (initems : (LA_Data.decomp * int) list) : bool list =
33245
65232054ffd0 eliminated some old folds;
wenzelm
parents: 33042
diff changeset
   739
  map fst (fold add_datoms initems []);
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   740
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   741
fun refutes ctxt params show_ex :
26945
9cd13e810998 renamed type decompT to decomp;
wenzelm
parents: 26835
diff changeset
   742
    (typ list * (LA_Data.decomp * int) list) list -> injust list -> injust list option =
9cd13e810998 renamed type decompT to decomp;
wenzelm
parents: 26835
diff changeset
   743
  let
9cd13e810998 renamed type decompT to decomp;
wenzelm
parents: 26835
diff changeset
   744
    fun refute ((Ts, initems : (LA_Data.decomp * int) list) :: initemss) (js: injust list) =
9cd13e810998 renamed type decompT to decomp;
wenzelm
parents: 26835
diff changeset
   745
          let
31510
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31102
diff changeset
   746
            val atoms = atoms_of (map fst initems)
26945
9cd13e810998 renamed type decompT to decomp;
wenzelm
parents: 26835
diff changeset
   747
            val n = length atoms
38052
04a8de29e8f7 dropped dead code
haftmann
parents: 36945
diff changeset
   748
            val mkleq = mklineq atoms
26945
9cd13e810998 renamed type decompT to decomp;
wenzelm
parents: 26835
diff changeset
   749
            val ixs = 0 upto (n - 1)
9cd13e810998 renamed type decompT to decomp;
wenzelm
parents: 26835
diff changeset
   750
            val iatoms = atoms ~~ ixs
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32740
diff changeset
   751
            val natlineqs = map_filter (mknat Ts ixs) iatoms
26945
9cd13e810998 renamed type decompT to decomp;
wenzelm
parents: 26835
diff changeset
   752
            val ineqs = map mkleq initems @ natlineqs
44654
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   753
          in case elim ctxt (ineqs, []) of
26945
9cd13e810998 renamed type decompT to decomp;
wenzelm
parents: 26835
diff changeset
   754
               Success j =>
44654
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   755
                 (trace_msg ctxt ("Contradiction! (" ^ string_of_int (length js + 1) ^ ")");
26945
9cd13e810998 renamed type decompT to decomp;
wenzelm
parents: 26835
diff changeset
   756
                  refute initemss (js @ [j]))
9cd13e810998 renamed type decompT to decomp;
wenzelm
parents: 26835
diff changeset
   757
             | Failure hist =>
44654
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   758
                 (if not show_ex orelse not (Config.get ctxt LA_Data.verbose) then ()
26945
9cd13e810998 renamed type decompT to decomp;
wenzelm
parents: 26835
diff changeset
   759
                  else
9cd13e810998 renamed type decompT to decomp;
wenzelm
parents: 26835
diff changeset
   760
                    let
9cd13e810998 renamed type decompT to decomp;
wenzelm
parents: 26835
diff changeset
   761
                      val (param_names, ctxt') = ctxt |> Variable.variant_fixes (map fst params)
9cd13e810998 renamed type decompT to decomp;
wenzelm
parents: 26835
diff changeset
   762
                      val (more_names, ctxt'') = ctxt' |> Variable.variant_fixes
43329
84472e198515 tuned signature: Name.invent and Name.invent_names;
wenzelm
parents: 42361
diff changeset
   763
                        (Name.invent (Variable.names_of ctxt') Name.uu (length Ts - length params))
26945
9cd13e810998 renamed type decompT to decomp;
wenzelm
parents: 26835
diff changeset
   764
                      val params' = (more_names @ param_names) ~~ Ts
9cd13e810998 renamed type decompT to decomp;
wenzelm
parents: 26835
diff changeset
   765
                    in
55362
5e5c36b051af disabled counterexample output for now; confusing because often incorrect
nipkow
parents: 52131
diff changeset
   766
                      () (*trace_ex ctxt'' params' atoms (discr initems) n hist*)
26945
9cd13e810998 renamed type decompT to decomp;
wenzelm
parents: 26835
diff changeset
   767
                    end; NONE)
9cd13e810998 renamed type decompT to decomp;
wenzelm
parents: 26835
diff changeset
   768
          end
9cd13e810998 renamed type decompT to decomp;
wenzelm
parents: 26835
diff changeset
   769
      | refute [] js = SOME js
9cd13e810998 renamed type decompT to decomp;
wenzelm
parents: 26835
diff changeset
   770
  in refute end;
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   771
30406
15dc25f8a0e2 Instead of giving up entirely, arith now ignores all inequalities when there are too many.
webertj
parents: 27020
diff changeset
   772
fun refute ctxt params show_ex do_pre split_neq terms : injust list option =
15dc25f8a0e2 Instead of giving up entirely, arith now ignores all inequalities when there are too many.
webertj
parents: 27020
diff changeset
   773
  refutes ctxt params show_ex (split_items ctxt do_pre split_neq
15dc25f8a0e2 Instead of giving up entirely, arith now ignores all inequalities when there are too many.
webertj
parents: 27020
diff changeset
   774
    (map snd params, terms)) [];
20254
58b71535ed00 lin_arith_prover splits certain operators (e.g. min, max, abs)
webertj
parents: 20217
diff changeset
   775
22950
haftmann
parents: 22846
diff changeset
   776
fun count P xs = length (filter P xs);
20254
58b71535ed00 lin_arith_prover splits certain operators (e.g. min, max, abs)
webertj
parents: 20217
diff changeset
   777
30406
15dc25f8a0e2 Instead of giving up entirely, arith now ignores all inequalities when there are too many.
webertj
parents: 27020
diff changeset
   778
fun prove ctxt params show_ex do_pre Hs concl : bool * injust list option =
20254
58b71535ed00 lin_arith_prover splits certain operators (e.g. min, max, abs)
webertj
parents: 20217
diff changeset
   779
  let
44654
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   780
    val _ = trace_msg ctxt "prove:"
20254
58b71535ed00 lin_arith_prover splits certain operators (e.g. min, max, abs)
webertj
parents: 20217
diff changeset
   781
    (* append the negated conclusion to 'Hs' -- this corresponds to     *)
58b71535ed00 lin_arith_prover splits certain operators (e.g. min, max, abs)
webertj
parents: 20217
diff changeset
   782
    (* '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: 20217
diff changeset
   783
    (* theorem/tactic level                                             *)
58b71535ed00 lin_arith_prover splits certain operators (e.g. min, max, abs)
webertj
parents: 20217
diff changeset
   784
    val Hs' = Hs @ [LA_Logic.neg_prop concl]
58b71535ed00 lin_arith_prover splits certain operators (e.g. min, max, abs)
webertj
parents: 20217
diff changeset
   785
    fun is_neq NONE                 = false
58b71535ed00 lin_arith_prover splits certain operators (e.g. min, max, abs)
webertj
parents: 20217
diff changeset
   786
      | is_neq (SOME (_,_,r,_,_,_)) = (r = "~=")
44654
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   787
    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: 27020
diff changeset
   788
    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: 20217
diff changeset
   789
  in
30406
15dc25f8a0e2 Instead of giving up entirely, arith now ignores all inequalities when there are too many.
webertj
parents: 27020
diff changeset
   790
    if split_neq then ()
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   791
    else
44654
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   792
      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: 27020
diff changeset
   793
        string_of_int neq_limit ^ "), ignoring all inequalities");
15dc25f8a0e2 Instead of giving up entirely, arith now ignores all inequalities when there are too many.
webertj
parents: 27020
diff changeset
   794
    (split_neq, refute ctxt params show_ex do_pre split_neq Hs')
23190
d45c4d6c5f15 fixed handling of meta-logic propositions
webertj
parents: 23063
diff changeset
   795
  end handle TERM ("neg_prop", _) =>
d45c4d6c5f15 fixed handling of meta-logic propositions
webertj
parents: 23063
diff changeset
   796
    (* since no meta-logic negation is available, we can only fail if   *)
d45c4d6c5f15 fixed handling of meta-logic propositions
webertj
parents: 23063
diff changeset
   797
    (* the conclusion is not of the form 'Trueprop $ _' (simply         *)
d45c4d6c5f15 fixed handling of meta-logic propositions
webertj
parents: 23063
diff changeset
   798
    (* dropping the conclusion doesn't work either, because even        *)
d45c4d6c5f15 fixed handling of meta-logic propositions
webertj
parents: 23063
diff changeset
   799
    (* 'False' does not imply arbitrary 'concl::prop')                  *)
44654
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   800
    (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: 27020
diff changeset
   801
      (false, NONE));
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   802
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49387
diff changeset
   803
fun refute_tac ctxt (i, split_neq, justs) =
6074
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   804
  fn state =>
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   805
    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: 31986
diff changeset
   806
      val _ = trace_thm ctxt
44654
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   807
        ["refute_tac (on subgoal " ^ string_of_int i ^ ", with " ^
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   808
          string_of_int (length justs) ^ " justification(s)):"] state
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   809
      val {neqE, ...} = get_data ctxt;
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   810
      fun just1 j =
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   811
        (* eliminate inequalities *)
30406
15dc25f8a0e2 Instead of giving up entirely, arith now ignores all inequalities when there are too many.
webertj
parents: 27020
diff changeset
   812
        (if split_neq then
15dc25f8a0e2 Instead of giving up entirely, arith now ignores all inequalities when there are too many.
webertj
parents: 27020
diff changeset
   813
          REPEAT_DETERM (eresolve_tac neqE i)
15dc25f8a0e2 Instead of giving up entirely, arith now ignores all inequalities when there are too many.
webertj
parents: 27020
diff changeset
   814
        else
15dc25f8a0e2 Instead of giving up entirely, arith now ignores all inequalities when there are too many.
webertj
parents: 27020
diff changeset
   815
          all_tac) THEN
44654
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   816
          PRIMITIVE (trace_thm ctxt ["State after neqE:"]) THEN
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   817
          (* use theorems generated from the actual justifications *)
58839
ccda99401bc8 eliminated aliases;
wenzelm
parents: 55362
diff changeset
   818
          Subgoal.FOCUS (fn {prems, ...} => resolve_tac [mkthm ctxt prems j] 1) ctxt i
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   819
    in
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   820
      (* rewrite "[| A1; ...; An |] ==> B" to "[| A1; ...; An; ~B |] ==> False" *)
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   821
      DETERM (resolve_tac [LA_Logic.notI, LA_Logic.ccontr] i) THEN
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   822
      (* user-defined preprocessing of the subgoal *)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49387
diff changeset
   823
      DETERM (LA_Data.pre_tac ctxt i) THEN
44654
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   824
      PRIMITIVE (trace_thm ctxt ["State after pre_tac:"]) THEN
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   825
      (* prove every resulting subgoal, using its justification *)
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   826
      EVERY (map just1 justs)
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   827
    end  state;
6074
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   828
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   829
(*
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   830
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
   831
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
   832
*)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49387
diff changeset
   833
fun simpset_lin_arith_tac ctxt show_ex = SUBGOAL (fn (A, i) =>
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   834
  let
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   835
    val params = rev (Logic.strip_params A)
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   836
    val Hs = Logic.strip_assums_hyp A
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   837
    val concl = Logic.strip_assums_concl A
44654
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   838
    val _ = trace_term ctxt ["Trying to refute subgoal " ^ string_of_int i] A
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   839
  in
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   840
    case prove ctxt params show_ex true Hs concl of
44654
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   841
      (_, NONE) => (trace_msg ctxt "Refutation failed."; no_tac)
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   842
    | (split_neq, SOME js) => (trace_msg ctxt "Refutation succeeded.";
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49387
diff changeset
   843
                               refute_tac ctxt (i, split_neq, js))
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   844
  end);
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   845
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49387
diff changeset
   846
fun prems_lin_arith_tac ctxt =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49387
diff changeset
   847
  Method.insert_tac (Simplifier.prems_of ctxt) THEN'
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49387
diff changeset
   848
  simpset_lin_arith_tac ctxt false;
17613
072c21e31b42 Simplifier.inherit_bounds;
wenzelm
parents: 17524
diff changeset
   849
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   850
fun lin_arith_tac ctxt =
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49387
diff changeset
   851
  simpset_lin_arith_tac (empty_simpset ctxt);
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   852
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   853
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   854
13186
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   855
(** Forward proof from theorems **)
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   856
20433
55471f940e5c lin_arith_prover: splitting reverted because of performance loss
webertj
parents: 20280
diff changeset
   857
(* 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: 20280
diff changeset
   858
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: 20280
diff changeset
   859
generated by function split_items. *)
55471f940e5c lin_arith_prover: splitting reverted because of performance loss
webertj
parents: 20280
diff changeset
   860
55471f940e5c lin_arith_prover: splitting reverted because of performance loss
webertj
parents: 20280
diff changeset
   861
datatype splittree = Tip of thm list
55471f940e5c lin_arith_prover: splitting reverted because of performance loss
webertj
parents: 20280
diff changeset
   862
                   | Spl of thm * cterm * splittree * cterm * splittree;
55471f940e5c lin_arith_prover: splitting reverted because of performance loss
webertj
parents: 20280
diff changeset
   863
55471f940e5c lin_arith_prover: splitting reverted because of performance loss
webertj
parents: 20280
diff changeset
   864
(* "(ct1 ==> ?R) ==> (ct2 ==> ?R) ==> ?R" is taken to (ct1, ct2) *)
55471f940e5c lin_arith_prover: splitting reverted because of performance loss
webertj
parents: 20280
diff changeset
   865
55471f940e5c lin_arith_prover: splitting reverted because of performance loss
webertj
parents: 20280
diff changeset
   866
fun extract (imp : cterm) : cterm * cterm =
55471f940e5c lin_arith_prover: splitting reverted because of performance loss
webertj
parents: 20280
diff changeset
   867
let val (Il, r)    = Thm.dest_comb imp
55471f940e5c lin_arith_prover: splitting reverted because of performance loss
webertj
parents: 20280
diff changeset
   868
    val (_, imp1)  = Thm.dest_comb Il
55471f940e5c lin_arith_prover: splitting reverted because of performance loss
webertj
parents: 20280
diff changeset
   869
    val (Ict1, _)  = Thm.dest_comb imp1
55471f940e5c lin_arith_prover: splitting reverted because of performance loss
webertj
parents: 20280
diff changeset
   870
    val (_, ct1)   = Thm.dest_comb Ict1
55471f940e5c lin_arith_prover: splitting reverted because of performance loss
webertj
parents: 20280
diff changeset
   871
    val (Ir, _)    = Thm.dest_comb r
55471f940e5c lin_arith_prover: splitting reverted because of performance loss
webertj
parents: 20280
diff changeset
   872
    val (_, Ict2r) = Thm.dest_comb Ir
55471f940e5c lin_arith_prover: splitting reverted because of performance loss
webertj
parents: 20280
diff changeset
   873
    val (Ict2, _)  = Thm.dest_comb Ict2r
55471f940e5c lin_arith_prover: splitting reverted because of performance loss
webertj
parents: 20280
diff changeset
   874
    val (_, ct2)   = Thm.dest_comb Ict2
55471f940e5c lin_arith_prover: splitting reverted because of performance loss
webertj
parents: 20280
diff changeset
   875
in (ct1, ct2) end;
55471f940e5c lin_arith_prover: splitting reverted because of performance loss
webertj
parents: 20280
diff changeset
   876
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   877
fun splitasms ctxt (asms : thm list) : splittree =
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   878
let val {neqE, ...} = get_data ctxt
35693
d58a4ac1ca1c Use same order of neq-elimination as in proof search.
hoelzl
parents: 35230
diff changeset
   879
    fun elim_neq [] (asms', []) = Tip (rev asms')
d58a4ac1ca1c Use same order of neq-elimination as in proof search.
hoelzl
parents: 35230
diff changeset
   880
      | elim_neq [] (asms', asms) = Tip (rev asms' @ asms)
49387
167708456269 dropped some unused identifiers
haftmann
parents: 47060
diff changeset
   881
      | elim_neq (_ :: neqs) (asms', []) = elim_neq neqs ([],rev asms')
35693
d58a4ac1ca1c Use same order of neq-elimination as in proof search.
hoelzl
parents: 35230
diff changeset
   882
      | elim_neq (neqs as (neq :: _)) (asms', asm::asms) =
d58a4ac1ca1c Use same order of neq-elimination as in proof search.
hoelzl
parents: 35230
diff changeset
   883
      (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: 20280
diff changeset
   884
        SOME spl =>
55471f940e5c lin_arith_prover: splitting reverted because of performance loss
webertj
parents: 20280
diff changeset
   885
          let val (ct1, ct2) = extract (cprop_of spl)
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   886
              val thm1 = Thm.assume ct1
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   887
              val thm2 = Thm.assume ct2
35693
d58a4ac1ca1c Use same order of neq-elimination as in proof search.
hoelzl
parents: 35230
diff changeset
   888
          in Spl (spl, ct1, elim_neq neqs (asms', asms@[thm1]),
d58a4ac1ca1c Use same order of neq-elimination as in proof search.
hoelzl
parents: 35230
diff changeset
   889
            ct2, elim_neq neqs (asms', asms@[thm2]))
20433
55471f940e5c lin_arith_prover: splitting reverted because of performance loss
webertj
parents: 20280
diff changeset
   890
          end
35693
d58a4ac1ca1c Use same order of neq-elimination as in proof search.
hoelzl
parents: 35230
diff changeset
   891
      | NONE => elim_neq neqs (asm::asms', asms))
d58a4ac1ca1c Use same order of neq-elimination as in proof search.
hoelzl
parents: 35230
diff changeset
   892
in elim_neq neqE ([], asms) end;
20433
55471f940e5c lin_arith_prover: splitting reverted because of performance loss
webertj
parents: 20280
diff changeset
   893
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49387
diff changeset
   894
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: 49387
diff changeset
   895
  | fwdproof ctxt (Spl (thm, ct1, tree1, ct2, tree2)) js =
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   896
      let
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49387
diff changeset
   897
        val (thm1, js1) = fwdproof ctxt tree1 js
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49387
diff changeset
   898
        val (thm2, js2) = fwdproof ctxt tree2 js1
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   899
        val thm1' = Thm.implies_intr ct1 thm1
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   900
        val thm2' = Thm.implies_intr ct2 thm2
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   901
      in (thm2' COMP (thm1' COMP thm), js2) end;
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   902
      (* FIXME needs handle THM _ => NONE ? *)
20433
55471f940e5c lin_arith_prover: splitting reverted because of performance loss
webertj
parents: 20280
diff changeset
   903
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49387
diff changeset
   904
fun prover ctxt thms Tconcl (js : injust list) split_neq pos : thm option =
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   905
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 41493
diff changeset
   906
    val thy = Proof_Context.theory_of ctxt
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   907
    val nTconcl = LA_Logic.neg_prop Tconcl
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   908
    val cnTconcl = cterm_of thy nTconcl
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   909
    val nTconclthm = Thm.assume cnTconcl
30406
15dc25f8a0e2 Instead of giving up entirely, arith now ignores all inequalities when there are too many.
webertj
parents: 27020
diff changeset
   910
    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: 49387
diff changeset
   911
    val (Falsethm, _) = fwdproof ctxt tree js
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   912
    val contr = if pos then LA_Logic.ccontr else LA_Logic.notI
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   913
    val concl = Thm.implies_intr cnTconcl Falsethm COMP contr
44654
d80fe56788a5 proper config option linarith_trace;
wenzelm
parents: 43609
diff changeset
   914
  in SOME (trace_thm ctxt ["Proved by lin. arith. prover:"] (LA_Logic.mk_Eq concl)) end
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   915
  (*in case concl contains ?-var, which makes assume fail:*)   (* FIXME Variable.import_terms *)
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   916
  handle THM _ => NONE;
13186
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   917
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   918
(* PRE: concl is not negated!
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   919
   This assumption is OK because
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   920
   1. lin_arith_simproc tries both to prove and disprove concl and
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   921
   2. lin_arith_simproc is applied by the Simplifier which
13186
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   922
      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: 13105
diff changeset
   923
*)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49387
diff changeset
   924
fun lin_arith_simproc ctxt concl =
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   925
  let
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49387
diff changeset
   926
    val thms = maps LA_Logic.atomize (Simplifier.prems_of ctxt)
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   927
    val Hs = map Thm.prop_of thms
6102
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
   928
    val Tconcl = LA_Logic.mk_Trueprop concl
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   929
  in
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   930
    case prove ctxt [] false false Hs Tconcl of (* concl provable? *)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49387
diff changeset
   931
      (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: 27020
diff changeset
   932
    | (_, NONE) =>
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   933
        let val nTconcl = LA_Logic.neg_prop Tconcl in
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   934
          case prove ctxt [] false false Hs nTconcl of (* ~concl provable? *)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49387
diff changeset
   935
            (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: 27020
diff changeset
   936
          | (_, NONE) => NONE
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   937
        end
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 24039
diff changeset
   938
  end;
6074
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   939
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   940
end;