| author | wenzelm | 
| Sat, 04 Feb 2017 19:10:12 +0100 | |
| changeset 64982 | c515464b4652 | 
| parent 64927 | a5a09855e424 | 
| child 66301 | 8a6a89d6cf2b | 
| permissions | -rw-r--r-- | 
| 63960 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1 | (* Title: HOL/ex/Argo_Examples.thy | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 2 | Author: Sascha Boehme | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 3 | *) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 4 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 5 | section \<open>Argo\<close> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 6 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 7 | theory Argo_Examples | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 8 | imports Complex_Main | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 9 | begin | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 10 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 11 | text \<open> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 12 | This theory is intended to showcase and test different features of the \<open>argo\<close> proof method. | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 13 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 14 | The \<open>argo\<close> proof method can be applied to propositional problems, problems involving equality | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 15 | reasoning and problems of linear real arithmetic. | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 16 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 17 | The \<open>argo\<close> proof method provides two options. To specify an upper limit of the proof methods | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 18 | run time in seconds, use the option \<open>argo_timeout\<close>. To specify the amount of output, use the | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 19 | option \<open>argo_trace\<close> with value \<open>none\<close> for no tracing output, value \<open>basic\<close> for viewing the | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 20 | underlying propositions and some timings, and value \<open>full\<close> for additionally inspecting the | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 21 | proof replay steps. | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 22 | \<close> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 23 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 24 | declare[[argo_trace = full]] | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 25 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 26 | subsection \<open>Propositional logic\<close> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 27 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 28 | notepad | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 29 | begin | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 30 | have "True" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 31 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 32 | have "~False" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 33 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 34 | fix P :: bool | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 35 | assume "False" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 36 | then have "P" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 37 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 38 | fix P :: bool | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 39 | assume "~True" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 40 | then have "P" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 41 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 42 | fix P :: bool | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 43 | assume "P" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 44 | then have "P" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 45 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 46 | fix P :: bool | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 47 | assume "~~P" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 48 | then have "P" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 49 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 50 | fix P Q R :: bool | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 51 | assume "P & Q & R" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 52 | then have "R & P & Q" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 53 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 54 | fix P Q R :: bool | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 55 | assume "P & (Q & True & R) & (Q & P) & True" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 56 | then have "R & P & Q" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 57 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 58 | fix P Q1 Q2 Q3 Q4 Q5 :: bool | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 59 | assume "Q1 & (Q2 & P & Q3) & (Q4 & ~P & Q5)" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 60 | then have "~True" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 61 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 62 | fix P Q1 Q2 Q3 :: bool | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 63 | assume "(Q1 & False) & Q2 & Q3" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 64 | then have "P::bool" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 65 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 66 | fix P Q R :: bool | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 67 | assume "P | Q | R" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 68 | then have "R | P | Q" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 69 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 70 | fix P Q R :: bool | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 71 | assume "P | (Q | False | R) | (Q | P) | False" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 72 | then have "R | P | Q" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 73 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 74 | fix P Q1 Q2 Q3 Q4 :: bool | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 75 | have "(Q1 & P & Q2) --> False | (Q3 | (Q4 | P) | False)" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 76 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 77 | fix Q1 Q2 Q3 Q4 :: bool | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 78 | have "Q1 | (Q2 | True | Q3) | Q4" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 79 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 80 | fix P Q R :: bool | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 81 | assume "(P & Q) | (P & ~Q) | (P & R) | (P & ~R)" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 82 | then have "P" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 83 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 84 | fix P :: bool | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 85 | assume "P = True" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 86 | then have "P" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 87 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 88 | fix P :: bool | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 89 | assume "False = P" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 90 | then have "~P" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 91 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 92 | fix P Q :: bool | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 93 | assume "P = (~P)" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 94 | then have "Q" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 95 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 96 | fix P :: bool | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 97 | have "(~P) = (~P)" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 98 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 99 | fix P Q :: bool | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 100 | assume "P" and "~Q" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 101 | then have "P = (~Q)" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 102 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 103 | fix P Q :: bool | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 104 | assume "((P::bool) = Q) | (Q = P)" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 105 | then have "(P --> Q) & (Q --> P)" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 106 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 107 | fix P Q :: bool | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 108 | assume "(P::bool) = Q" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 109 | then have "Q = P" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 110 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 111 | fix P Q R :: bool | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 112 | assume "if P then Q else R" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 113 | then have "Q | R" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 114 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 115 | fix P Q :: bool | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 116 | assume "P | Q" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 117 | and "P | ~Q" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 118 | and "~P | Q" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 119 | and "~P | ~Q" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 120 | then have "False" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 121 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 122 | fix P Q R :: bool | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 123 | assume "P | Q | R" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 124 | and "P | Q | ~R" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 125 | and "P | ~Q | R" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 126 | and "P | ~Q | ~R" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 127 | and "~P | Q | R" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 128 | and "~P | Q | ~R" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 129 | and "~P | ~Q | R" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 130 | and "~P | ~Q | ~R" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 131 | then have "False" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 132 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 133 | fix a b c d e f g h i j k l m n p q :: bool | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 134 | assume "(a & b | c & d) & (e & f | g & h) | (i & j | k & l) & (m & n | p & q)" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 135 | then have "(a & b | c & d) & (e & f | g & h) | (i & j | k & l) & (m & n | p & q)" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 136 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 137 | fix P :: bool | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 138 | have "P=P=P=P=P=P=P=P=P=P" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 139 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 140 | fix a b c d e f p q x :: bool | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 141 | assume "a | b | c | d" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 142 | and "e | f | (a & d)" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 143 | and "~(a | (c & ~c)) | b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 144 | and "~(b & (x | ~x)) | c" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 145 | and "~(d | False) | c" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 146 | and "~(c | (~p & (p | (q & ~q))))" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 147 | then have "False" by argo | 
| 64927 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63960diff
changeset | 148 | next | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63960diff
changeset | 149 | have "(True & True & True) = True" by argo | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63960diff
changeset | 150 | next | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63960diff
changeset | 151 | have "(False | False | False) = False" by argo | 
| 63960 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 152 | end | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 153 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 154 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 155 | subsection \<open>Equality, congruence and predicates\<close> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 156 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 157 | notepad | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 158 | begin | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 159 | fix t :: "'a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 160 | have "t = t" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 161 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 162 | fix t u :: "'a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 163 | assume "t = u" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 164 | then have "u = t" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 165 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 166 | fix s t u :: "'a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 167 | assume "s = t" and "t = u" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 168 | then have "s = u" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 169 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 170 | fix s t u v :: "'a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 171 | assume "s = t" and "t = u" and "u = v" and "u = s" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 172 | then have "s = v" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 173 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 174 | fix s t u v w :: "'a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 175 | assume "s = t" and "t = u" and "s = v" and "v = w" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 176 | then have "w = u" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 177 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 178 | fix s t u a b c :: "'a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 179 | assume "s = t" and "t = u" and "a = b" and "b = c" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 180 | then have "s = a --> c = u" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 181 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 182 | fix a b c d :: "'a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 183 | assume "(a = b & b = c) | (a = d & d = c)" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 184 | then have "a = c" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 185 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 186 | fix a b1 b2 b3 b4 c d :: "'a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 187 | assume "(a = b1 & ((b1 = b2 & b2 = b3) | (b1 = b4 & b4 = b3)) & b3 = c) | (a = d & d = c)" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 188 | then have "a = c" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 189 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 190 | fix a b :: "'a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 191 | have "(if True then a else b) = a" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 192 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 193 | fix a b :: "'a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 194 | have "(if False then a else b) = b" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 195 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 196 | fix a b :: "'a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 197 | have "(if \<not>True then a else b) = b" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 198 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 199 | fix a b :: "'a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 200 | have "(if \<not>False then a else b) = a" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 201 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 202 | fix P :: "bool" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 203 | fix a :: "'a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 204 | have "(if P then a else a) = a" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 205 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 206 | fix P :: "bool" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 207 | fix a b c :: "'a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 208 | assume "P" and "a = c" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 209 | then have "(if P then a else b) = c" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 210 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 211 | fix P :: "bool" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 212 | fix a b c :: "'a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 213 | assume "~P" and "b = c" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 214 | then have "(if P then a else b) = c" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 215 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 216 | fix P Q :: "bool" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 217 | fix a b c d :: "'a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 218 | assume "P" and "Q" and "a = d" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 219 | then have "(if P then (if Q then a else b) else c) = d" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 220 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 221 | fix a b c :: "'a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 222 | assume "a \<noteq> b" and "b = c" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 223 | then have "a \<noteq> c" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 224 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 225 | fix a b c :: "'a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 226 | assume "a \<noteq> b" and "a = c" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 227 | then have "c \<noteq> b" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 228 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 229 | fix a b c d :: "'a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 230 | assume "a = b" and "c = d" and "b \<noteq> d" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 231 | then have "a \<noteq> c" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 232 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 233 | fix a b c d :: "'a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 234 | assume "a = b" and "c = d" and "d \<noteq> b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 235 | then have "a \<noteq> c" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 236 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 237 | fix a b c d :: "'a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 238 | assume "a = b" and "c = d" and "b \<noteq> d" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 239 | then have "c \<noteq> a" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 240 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 241 | fix a b c d :: "'a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 242 | assume "a = b" and "c = d" and "d \<noteq> b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 243 | then have "c \<noteq> a" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 244 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 245 | fix a b c d e f :: "'a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 246 | assume "a \<noteq> b" and "b = c" and "b = d" and "d = e" and "a = f" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 247 | then have "f \<noteq> e" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 248 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 249 | fix a b :: "'a" and f :: "'a \<Rightarrow> 'a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 250 | assume "a = b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 251 | then have "f a = f b" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 252 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 253 | fix a b c :: "'a" and f :: "'a \<Rightarrow> 'a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 254 | assume "f a = f b" and "b = c" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 255 | then have "f a = f c" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 256 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 257 | fix a :: "'a" and f :: "'a \<Rightarrow> 'a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 258 | assume "f a = a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 259 | then have "f (f a) = a" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 260 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 261 | fix a b :: "'a" and f g :: "'a \<Rightarrow> 'a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 262 | assume "a = b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 263 | then have "g (f a) = g (f b)" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 264 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 265 | fix a b :: "'a" and f g :: "'a \<Rightarrow> 'a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 266 | assume "f a = b" and "g b = a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 267 | then have "f (g (f a)) = b" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 268 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 269 | fix a b :: "'a" and g :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 270 | assume "a = b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 271 | then have "g a b = g b a" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 272 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 273 | fix a b :: "'a" and f :: "'a \<Rightarrow> 'a" and g :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 274 | assume "f a = b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 275 | then have "g (f a) b = g b (f a)" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 276 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 277 | fix a b c d e g h :: "'a" and f :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 278 | assume "c = d" and "e = c" and "e = b" and "b = h" and "f g h = d" and "f g d = a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 279 | then have "a = b" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 280 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 281 | fix a b :: "'a" and P :: "'a \<Rightarrow> bool" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 282 | assume "P a" and "a = b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 283 | then have "P b" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 284 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 285 | fix a b :: "'a" and P :: "'a \<Rightarrow> bool" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 286 | assume "~ P a" and "a = b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 287 | then have "~ P b" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 288 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 289 | fix a b c d :: "'a" and P :: "'a \<Rightarrow> 'a \<Rightarrow> bool" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 290 | assume "P a b" and "a = c" and "b = d" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 291 | then have "P c d" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 292 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 293 | fix a b c d :: "'a" and P :: "'a \<Rightarrow> 'a \<Rightarrow> bool" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 294 | assume "~ P a b" and "a = c" and "b = d" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 295 | then have "~ P c d" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 296 | end | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 297 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 298 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 299 | subsection \<open>Linear real arithmetic\<close> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 300 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 301 | subsubsection \<open>Negation and subtraction\<close> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 302 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 303 | notepad | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 304 | begin | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 305 | fix a b :: real | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 306 | have | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 307 | "-a = -1 * a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 308 | "-(-a) = a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 309 | "a - b = a + -1 * b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 310 | "a - (-b) = a + b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 311 | by argo+ | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 312 | end | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 313 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 314 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 315 | subsubsection \<open>Multiplication\<close> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 316 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 317 | notepad | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 318 | begin | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 319 | fix a b c d :: real | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 320 | have | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 321 | "(2::real) * 3 = 6" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 322 | "0 * a = 0" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 323 | "a * 0 = 0" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 324 | "1 * a = a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 325 | "a * 1 = a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 326 | "2 * a = a * 2" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 327 | "2 * a * 3 = 6 * a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 328 | "2 * a * 3 * 5 = 30 * a" | 
| 64927 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63960diff
changeset | 329 | "2 * (a * (3 * 5)) = 30 * a" | 
| 63960 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 330 | "a * 0 * b = 0" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 331 | "a * (0 * b) = 0" | 
| 64927 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63960diff
changeset | 332 | "a * b = b * a" | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63960diff
changeset | 333 | "a * b * a = b * a * a" | 
| 63960 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 334 | "a * (b * c) = (a * b) * c" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 335 | "a * (b * (c * d)) = ((a * b) * c) * d" | 
| 64927 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63960diff
changeset | 336 | "a * (b * (c * d)) = ((d * c) * b) * a" | 
| 63960 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 337 | "a * (b + c + d) = a * b + a * c + a * d" | 
| 64927 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63960diff
changeset | 338 | "(a + b + c) * d = a * d + b * d + c * d" | 
| 63960 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 339 | by argo+ | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 340 | end | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 341 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 342 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 343 | subsubsection \<open>Division\<close> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 344 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 345 | notepad | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 346 | begin | 
| 64927 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63960diff
changeset | 347 | fix a b c d :: real | 
| 63960 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 348 | have | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 349 | "(6::real) / 2 = 3" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 350 | "a / 0 = a / 0" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 351 | "a / 0 <= a / 0" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 352 | "~(a / 0 < a / 0)" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 353 | "0 / a = 0" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 354 | "a / 1 = a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 355 | "a / 3 = 1/3 * a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 356 | "6 * a / 2 = 3 * a" | 
| 64927 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63960diff
changeset | 357 | "(6 * a) / 2 = 3 * a" | 
| 63960 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 358 | "a / ((5 * b) / 2) = 2/5 * a / b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 359 | "a / (5 * (b / 2)) = 2/5 * a / b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 360 | "(a / 5) * (b / 2) = 1/10 * a * b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 361 | "a / (3 * b) = 1/3 * a / b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 362 | "(a + b) / 5 = 1/5 * a + 1/5 * b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 363 | "a / (5 * 1/5) = a" | 
| 64927 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63960diff
changeset | 364 | "a * (b / c) = (b * a) / c" | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63960diff
changeset | 365 | "(a / b) * c = (c * a) / b" | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63960diff
changeset | 366 | "(a / b) / (c / d) = (a * d) / (c * b)" | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63960diff
changeset | 367 | "1 / (a * b) = 1 / (b * a)" | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63960diff
changeset | 368 | "a / (3 * b) = 1 / 3 * a / b" | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63960diff
changeset | 369 | "(a + b + c) / d = a / d + b / d + c / d" | 
| 63960 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 370 | by argo+ | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 371 | end | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 372 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 373 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 374 | subsubsection \<open>Addition\<close> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 375 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 376 | notepad | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 377 | begin | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 378 | fix a b c d :: real | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 379 | have | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 380 | "a + b = b + a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 381 | "a + b + c = c + b + a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 382 | "a + b + c + d = d + c + b + a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 383 | "a + (b + (c + d)) = ((a + b) + c) + d" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 384 | "(5::real) + -3 = 2" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 385 | "(3::real) + 5 + -1 = 7" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 386 | "2 + a = a + 2" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 387 | "a + b + a = b + 2 * a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 388 | "-1 + a + -1 + 2 + b + 5/3 + b + 1/3 + 5 * b + 2/3 = 8/3 + a + 7 * b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 389 | "1 + b + b + 5 * b + 3 * a + 7 + a + 2 = 10 + 4 * a + 7 * b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 390 | by argo+ | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 391 | end | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 392 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 393 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 394 | subsubsection \<open>Minimum and maximum\<close> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 395 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 396 | notepad | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 397 | begin | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 398 | fix a b :: real | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 399 | have | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 400 | "min (3::real) 5 = 3" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 401 | "min (5::real) 3 = 3" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 402 | "min (3::real) (-5) = -5" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 403 | "min (-5::real) 3 = -5" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 404 | "min a a = a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 405 | "a \<le> b \<longrightarrow> min a b = a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 406 | "a > b \<longrightarrow> min a b = b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 407 | "min a b \<le> a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 408 | "min a b \<le> b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 409 | "min a b = min b a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 410 | by argo+ | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 411 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 412 | fix a b :: real | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 413 | have | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 414 | "max (3::real) 5 = 5" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 415 | "max (5::real) 3 = 5" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 416 | "max (3::real) (-5) = 3" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 417 | "max (-5::real) 3 = 3" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 418 | "max a a = a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 419 | "a \<le> b \<longrightarrow> max a b = b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 420 | "a > b \<longrightarrow> max a b = a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 421 | "a \<le> max a b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 422 | "b \<le> max a b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 423 | "max a b = max b a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 424 | by argo+ | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 425 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 426 | fix a b :: real | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 427 | have | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 428 | "min a b \<le> max a b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 429 | "min a b + max a b = a + b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 430 | "a < b \<longrightarrow> min a b < max a b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 431 | by argo+ | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 432 | end | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 433 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 434 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 435 | subsubsection \<open>Absolute value\<close> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 436 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 437 | notepad | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 438 | begin | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 439 | fix a :: real | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 440 | have | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 441 | "abs (3::real) = 3" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 442 | "abs (-3::real) = 3" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 443 | "0 \<le> abs a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 444 | "a \<le> abs a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 445 | "a \<ge> 0 \<longrightarrow> abs a = a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 446 | "a < 0 \<longrightarrow> abs a = -a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 447 | "abs (abs a) = abs a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 448 | by argo+ | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 449 | end | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 450 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 451 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 452 | subsubsection \<open>Equality\<close> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 453 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 454 | notepad | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 455 | begin | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 456 | fix a b c d :: real | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 457 | have | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 458 | "(3::real) = 3" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 459 | "~((3::real) = 4)" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 460 | "~((4::real) = 3)" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 461 | "3 * a = 5 --> a = 5/3" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 462 | "-3 * a = 5 --> -5/3 = a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 463 | "5 = 3 * a --> 5/3 = a " | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 464 | "5 = -3 * a --> a = -5/3" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 465 | "2 + 3 * a = 4 --> a = 2/3" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 466 | "4 = 2 + 3 * a --> 2/3 = a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 467 | "2 + 3 * a + 5 * b + c = 4 --> 3 * a + 5 * b + c = 2" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 468 | "4 = 2 + 3 * a + 5 * b + c --> 2 = 3 * a + 5 * b + c" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 469 | "-2 * a + b + -3 * c = 7 --> -7 = 2 * a + -1 * b + 3 * c" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 470 | "7 = -2 * a + b + -3 * c --> 2 * a + -1 * b + 3 * c = -7" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 471 | "-2 * a + b + -3 * c + 4 * d = 7 --> -7 = 2 * a + -1 * b + 3 * c + -4 * d" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 472 | "7 = -2 * a + b + -3 * c + 4 * d --> 2 * a + -1 * b + 3 * c + -4 * d = -7" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 473 | "a + 3 * b = 5 * c + b --> a + 2 * b + -5 * c = 0" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 474 | by argo+ | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 475 | end | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 476 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 477 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 478 | subsubsection \<open>Less-equal\<close> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 479 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 480 | notepad | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 481 | begin | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 482 | fix a b c d :: real | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 483 | have | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 484 | "(3::real) <= 3" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 485 | "(3::real) <= 4" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 486 | "~((4::real) <= 3)" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 487 | "3 * a <= 5 --> a <= 5/3" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 488 | "-3 * a <= 5 --> -5/3 <= a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 489 | "5 <= 3 * a --> 5/3 <= a " | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 490 | "5 <= -3 * a --> a <= -5/3" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 491 | "2 + 3 * a <= 4 --> a <= 2/3" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 492 | "4 <= 2 + 3 * a --> 2/3 <= a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 493 | "2 + 3 * a + 5 * b + c <= 4 --> 3 * a + 5 * b + c <= 2" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 494 | "4 <= 2 + 3 * a + 5 * b + c --> 2 <= 3 * a + 5 * b + c" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 495 | "-2 * a + b + -3 * c <= 7 --> -7 <= 2 * a + -1 * b + 3 * c" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 496 | "7 <= -2 * a + b + -3 * c --> 2 * a + -1 * b + 3 * c <= -7" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 497 | "-2 * a + b + -3 * c + 4 * d <= 7 --> -7 <= 2 * a + -1 * b + 3 * c + -4 * d" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 498 | "7 <= -2 * a + b + -3 * c + 4 * d --> 2 * a + -1 * b + 3 * c + -4 * d <= -7" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 499 | "a + 3 * b <= 5 * c + b --> a + 2 * b + -5 * c <= 0" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 500 | by argo+ | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 501 | end | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 502 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 503 | subsubsection \<open>Less\<close> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 504 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 505 | notepad | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 506 | begin | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 507 | fix a b c d :: real | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 508 | have | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 509 | "(3::real) < 4" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 510 | "~((3::real) < 3)" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 511 | "~((4::real) < 3)" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 512 | "3 * a < 5 --> a < 5/3" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 513 | "-3 * a < 5 --> -5/3 < a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 514 | "5 < 3 * a --> 5/3 < a " | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 515 | "5 < -3 * a --> a < -5/3" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 516 | "2 + 3 * a < 4 --> a < 2/3" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 517 | "4 < 2 + 3 * a --> 2/3 < a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 518 | "2 + 3 * a + 5 * b + c < 4 --> 3 * a + 5 * b + c < 2" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 519 | "4 < 2 + 3 * a + 5 * b + c --> 2 < 3 * a + 5 * b + c" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 520 | "-2 * a + b + -3 * c < 7 --> -7 < 2 * a + -1 * b + 3 * c" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 521 | "7 < -2 * a + b + -3 * c --> 2 * a + -1 * b + 3 * c < -7" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 522 | "-2 * a + b + -3 * c + 4 * d < 7 --> -7 < 2 * a + -1 * b + 3 * c + -4 * d" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 523 | "7 < -2 * a + b + -3 * c + 4 * d --> 2 * a + -1 * b + 3 * c + -4 * d < -7" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 524 | "a + 3 * b < 5 * c + b --> a + 2 * b + -5 * c < 0" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 525 | by argo+ | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 526 | end | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 527 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 528 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 529 | subsubsection \<open>Other examples\<close> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 530 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 531 | notepad | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 532 | begin | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 533 | have | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 534 | "(0::real) < 1" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 535 | "(47::real) + 11 < 8 * 15" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 536 | by argo+ | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 537 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 538 | fix a :: real | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 539 | assume "a < 3" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 540 | then have "a < 5" "a <= 5" "~(5 < a)" "~(5 <= a)" by argo+ | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 541 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 542 | fix a :: real | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 543 | assume "a <= 3" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 544 | then have "a < 5" "a <= 5" "~(5 < a)" "~(5 <= a)" by argo+ | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 545 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 546 | fix a :: real | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 547 | assume "~(3 < a)" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 548 | then have "a < 5" "a <= 5" "~(5 < a)" "~(5 <= a)" by argo+ | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 549 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 550 | fix a :: real | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 551 | assume "~(3 <= a)" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 552 | then have "a < 5" "a <= 5" "~(5 < a)" "~(5 <= a)" by argo+ | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 553 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 554 | fix a :: real | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 555 | have "a < 3 | a = 3 | a > 3" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 556 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 557 | fix a b :: real | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 558 | assume "0 < a" and "a < b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 559 | then have "0 < b" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 560 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 561 | fix a b :: real | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 562 | assume "0 < a" and "a \<le> b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 563 | then have "0 \<le> b" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 564 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 565 | fix a b :: real | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 566 | assume "0 \<le> a" and "a < b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 567 | then have "0 \<le> b" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 568 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 569 | fix a b :: real | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 570 | assume "0 \<le> a" and "a \<le> b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 571 | then have "0 \<le> b" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 572 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 573 | fix a b c :: real | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 574 | assume "2 \<le> a" and "3 \<le> b" and "c \<le> 5" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 575 | then have "-2 * a + -3 * b + 5 * c < 13" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 576 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 577 | fix a b c :: real | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 578 | assume "2 \<le> a" and "3 \<le> b" and "c \<le> 5" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 579 | then have "-2 * a + -3 * b + 5 * c \<le> 12" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 580 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 581 | fix a b :: real | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 582 | assume "a = 2" and "b = 3" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 583 | then have "a + b > 5 \<or> a < b" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 584 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 585 | fix a b c :: real | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 586 | assume "5 < b + c" and "a + c < 0" and "a > 0" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 587 | then have "b > 0" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 588 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 589 | fix a b c :: real | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 590 | assume "a + b < 7" and "5 < b + c" and "a + c < 0" and "a > 0" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 591 | then have "0 < b \<and> b < 7" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 592 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 593 | fix a b c :: real | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 594 | assume "a < b" and "b < c" and "c < a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 595 | then have "False" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 596 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 597 | fix a b :: real | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 598 | assume "a - 5 > b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 599 | then have "b < a" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 600 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 601 | fix a b :: real | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 602 | have "(a - b) - a = (a - a) - b" by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 603 | next | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 604 | fix n m n' m' :: real | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 605 | have " | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 606 | (n < m & m < n') | (n < m & m = n') | (n < n' & n' < m) | | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 607 | (n = n' & n' < m) | (n = m & m < n') | | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 608 | (n' < m & m < n) | (n' < m & m = n) | | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 609 | (n' < n & n < m) | (n' = n & n < m) | (n' = m & m < n) | | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 610 | (m < n & n < n') | (m < n & n' = n) | (m < n' & n' < n) | | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 611 | (m = n & n < n') | (m = n' & n' < n) | | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 612 | (n' = m & m = n)" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 613 | by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 614 | end | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 615 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 616 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 617 | subsection \<open>Larger examples\<close> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 618 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 619 | declare[[argo_trace = basic, argo_timeout = 60]] | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 620 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 621 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 622 | text \<open>Translated from TPTP problem library: PUZ015-2.006.dimacs\<close> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 623 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 624 | lemma assumes 1: "~x0" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 625 | and 2: "~x30" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 626 | and 3: "~x29" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 627 | and 4: "~x59" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 628 | and 5: "x1 | x31 | x0" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 629 | and 6: "x2 | x32 | x1" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 630 | and 7: "x3 | x33 | x2" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 631 | and 8: "x4 | x34 | x3" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 632 | and 9: "x35 | x4" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 633 | and 10: "x5 | x36 | x30" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 634 | and 11: "x6 | x37 | x5 | x31" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 635 | and 12: "x7 | x38 | x6 | x32" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 636 | and 13: "x8 | x39 | x7 | x33" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 637 | and 14: "x9 | x40 | x8 | x34" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 638 | and 15: "x41 | x9 | x35" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 639 | and 16: "x10 | x42 | x36" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 640 | and 17: "x11 | x43 | x10 | x37" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 641 | and 18: "x12 | x44 | x11 | x38" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 642 | and 19: "x13 | x45 | x12 | x39" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 643 | and 20: "x14 | x46 | x13 | x40" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 644 | and 21: "x47 | x14 | x41" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 645 | and 22: "x15 | x48 | x42" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 646 | and 23: "x16 | x49 | x15 | x43" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 647 | and 24: "x17 | x50 | x16 | x44" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 648 | and 25: "x18 | x51 | x17 | x45" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 649 | and 26: "x19 | x52 | x18 | x46" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 650 | and 27: "x53 | x19 | x47" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 651 | and 28: "x20 | x54 | x48" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 652 | and 29: "x21 | x55 | x20 | x49" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 653 | and 30: "x22 | x56 | x21 | x50" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 654 | and 31: "x23 | x57 | x22 | x51" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 655 | and 32: "x24 | x58 | x23 | x52" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 656 | and 33: "x59 | x24 | x53" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 657 | and 34: "x25 | x54" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 658 | and 35: "x26 | x25 | x55" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 659 | and 36: "x27 | x26 | x56" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 660 | and 37: "x28 | x27 | x57" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 661 | and 38: "x29 | x28 | x58" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 662 | and 39: "~x1 | ~x31" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 663 | and 40: "~x1 | ~x0" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 664 | and 41: "~x31 | ~x0" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 665 | and 42: "~x2 | ~x32" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 666 | and 43: "~x2 | ~x1" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 667 | and 44: "~x32 | ~x1" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 668 | and 45: "~x3 | ~x33" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 669 | and 46: "~x3 | ~x2" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 670 | and 47: "~x33 | ~x2" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 671 | and 48: "~x4 | ~x34" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 672 | and 49: "~x4 | ~x3" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 673 | and 50: "~x34 | ~x3" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 674 | and 51: "~x35 | ~x4" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 675 | and 52: "~x5 | ~x36" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 676 | and 53: "~x5 | ~x30" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 677 | and 54: "~x36 | ~x30" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 678 | and 55: "~x6 | ~x37" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 679 | and 56: "~x6 | ~x5" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 680 | and 57: "~x6 | ~x31" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 681 | and 58: "~x37 | ~x5" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 682 | and 59: "~x37 | ~x31" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 683 | and 60: "~x5 | ~x31" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 684 | and 61: "~x7 | ~x38" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 685 | and 62: "~x7 | ~x6" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 686 | and 63: "~x7 | ~x32" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 687 | and 64: "~x38 | ~x6" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 688 | and 65: "~x38 | ~x32" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 689 | and 66: "~x6 | ~x32" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 690 | and 67: "~x8 | ~x39" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 691 | and 68: "~x8 | ~x7" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 692 | and 69: "~x8 | ~x33" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 693 | and 70: "~x39 | ~x7" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 694 | and 71: "~x39 | ~x33" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 695 | and 72: "~x7 | ~x33" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 696 | and 73: "~x9 | ~x40" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 697 | and 74: "~x9 | ~x8" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 698 | and 75: "~x9 | ~x34" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 699 | and 76: "~x40 | ~x8" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 700 | and 77: "~x40 | ~x34" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 701 | and 78: "~x8 | ~x34" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 702 | and 79: "~x41 | ~x9" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 703 | and 80: "~x41 | ~x35" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 704 | and 81: "~x9 | ~x35" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 705 | and 82: "~x10 | ~x42" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 706 | and 83: "~x10 | ~x36" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 707 | and 84: "~x42 | ~x36" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 708 | and 85: "~x11 | ~x43" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 709 | and 86: "~x11 | ~x10" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 710 | and 87: "~x11 | ~x37" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 711 | and 88: "~x43 | ~x10" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 712 | and 89: "~x43 | ~x37" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 713 | and 90: "~x10 | ~x37" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 714 | and 91: "~x12 | ~x44" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 715 | and 92: "~x12 | ~x11" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 716 | and 93: "~x12 | ~x38" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 717 | and 94: "~x44 | ~x11" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 718 | and 95: "~x44 | ~x38" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 719 | and 96: "~x11 | ~x38" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 720 | and 97: "~x13 | ~x45" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 721 | and 98: "~x13 | ~x12" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 722 | and 99: "~x13 | ~x39" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 723 | and 100: "~x45 | ~x12" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 724 | and 101: "~x45 | ~x39" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 725 | and 102: "~x12 | ~x39" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 726 | and 103: "~x14 | ~x46" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 727 | and 104: "~x14 | ~x13" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 728 | and 105: "~x14 | ~x40" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 729 | and 106: "~x46 | ~x13" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 730 | and 107: "~x46 | ~x40" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 731 | and 108: "~x13 | ~x40" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 732 | and 109: "~x47 | ~x14" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 733 | and 110: "~x47 | ~x41" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 734 | and 111: "~x14 | ~x41" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 735 | and 112: "~x15 | ~x48" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 736 | and 113: "~x15 | ~x42" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 737 | and 114: "~x48 | ~x42" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 738 | and 115: "~x16 | ~x49" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 739 | and 116: "~x16 | ~x15" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 740 | and 117: "~x16 | ~x43" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 741 | and 118: "~x49 | ~x15" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 742 | and 119: "~x49 | ~x43" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 743 | and 120: "~x15 | ~x43" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 744 | and 121: "~x17 | ~x50" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 745 | and 122: "~x17 | ~x16" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 746 | and 123: "~x17 | ~x44" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 747 | and 124: "~x50 | ~x16" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 748 | and 125: "~x50 | ~x44" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 749 | and 126: "~x16 | ~x44" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 750 | and 127: "~x18 | ~x51" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 751 | and 128: "~x18 | ~x17" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 752 | and 129: "~x18 | ~x45" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 753 | and 130: "~x51 | ~x17" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 754 | and 131: "~x51 | ~x45" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 755 | and 132: "~x17 | ~x45" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 756 | and 133: "~x19 | ~x52" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 757 | and 134: "~x19 | ~x18" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 758 | and 135: "~x19 | ~x46" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 759 | and 136: "~x52 | ~x18" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 760 | and 137: "~x52 | ~x46" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 761 | and 138: "~x18 | ~x46" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 762 | and 139: "~x53 | ~x19" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 763 | and 140: "~x53 | ~x47" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 764 | and 141: "~x19 | ~x47" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 765 | and 142: "~x20 | ~x54" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 766 | and 143: "~x20 | ~x48" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 767 | and 144: "~x54 | ~x48" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 768 | and 145: "~x21 | ~x55" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 769 | and 146: "~x21 | ~x20" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 770 | and 147: "~x21 | ~x49" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 771 | and 148: "~x55 | ~x20" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 772 | and 149: "~x55 | ~x49" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 773 | and 150: "~x20 | ~x49" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 774 | and 151: "~x22 | ~x56" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 775 | and 152: "~x22 | ~x21" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 776 | and 153: "~x22 | ~x50" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 777 | and 154: "~x56 | ~x21" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 778 | and 155: "~x56 | ~x50" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 779 | and 156: "~x21 | ~x50" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 780 | and 157: "~x23 | ~x57" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 781 | and 158: "~x23 | ~x22" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 782 | and 159: "~x23 | ~x51" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 783 | and 160: "~x57 | ~x22" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 784 | and 161: "~x57 | ~x51" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 785 | and 162: "~x22 | ~x51" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 786 | and 163: "~x24 | ~x58" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 787 | and 164: "~x24 | ~x23" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 788 | and 165: "~x24 | ~x52" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 789 | and 166: "~x58 | ~x23" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 790 | and 167: "~x58 | ~x52" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 791 | and 168: "~x23 | ~x52" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 792 | and 169: "~x59 | ~x24" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 793 | and 170: "~x59 | ~x53" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 794 | and 171: "~x24 | ~x53" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 795 | and 172: "~x25 | ~x54" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 796 | and 173: "~x26 | ~x25" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 797 | and 174: "~x26 | ~x55" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 798 | and 175: "~x25 | ~x55" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 799 | and 176: "~x27 | ~x26" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 800 | and 177: "~x27 | ~x56" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 801 | and 178: "~x26 | ~x56" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 802 | and 179: "~x28 | ~x27" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 803 | and 180: "~x28 | ~x57" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 804 | and 181: "~x27 | ~x57" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 805 | and 182: "~x29 | ~x28" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 806 | and 183: "~x29 | ~x58" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 807 | and 184: "~x28 | ~x58" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 808 | shows "False" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 809 | using assms | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 810 | by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 811 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 812 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 813 | text \<open>Translated from TPTP problem library: MSC007-1.008.dimacs\<close> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 814 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 815 | lemma assumes 1: "x0 | x1 | x2 | x3 | x4 | x5 | x6" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 816 | and 2: "x7 | x8 | x9 | x10 | x11 | x12 | x13" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 817 | and 3: "x14 | x15 | x16 | x17 | x18 | x19 | x20" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 818 | and 4: "x21 | x22 | x23 | x24 | x25 | x26 | x27" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 819 | and 5: "x28 | x29 | x30 | x31 | x32 | x33 | x34" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 820 | and 6: "x35 | x36 | x37 | x38 | x39 | x40 | x41" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 821 | and 7: "x42 | x43 | x44 | x45 | x46 | x47 | x48" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 822 | and 8: "x49 | x50 | x51 | x52 | x53 | x54 | x55" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 823 | and 9: "~x0 | ~x7" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 824 | and 10: "~x0 | ~x14" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 825 | and 11: "~x0 | ~x21" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 826 | and 12: "~x0 | ~x28" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 827 | and 13: "~x0 | ~x35" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 828 | and 14: "~x0 | ~x42" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 829 | and 15: "~x0 | ~x49" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 830 | and 16: "~x7 | ~x14" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 831 | and 17: "~x7 | ~x21" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 832 | and 18: "~x7 | ~x28" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 833 | and 19: "~x7 | ~x35" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 834 | and 20: "~x7 | ~x42" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 835 | and 21: "~x7 | ~x49" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 836 | and 22: "~x14 | ~x21" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 837 | and 23: "~x14 | ~x28" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 838 | and 24: "~x14 | ~x35" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 839 | and 25: "~x14 | ~x42" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 840 | and 26: "~x14 | ~x49" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 841 | and 27: "~x21 | ~x28" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 842 | and 28: "~x21 | ~x35" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 843 | and 29: "~x21 | ~x42" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 844 | and 30: "~x21 | ~x49" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 845 | and 31: "~x28 | ~x35" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 846 | and 32: "~x28 | ~x42" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 847 | and 33: "~x28 | ~x49" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 848 | and 34: "~x35 | ~x42" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 849 | and 35: "~x35 | ~x49" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 850 | and 36: "~x42 | ~x49" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 851 | and 37: "~x1 | ~x8" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 852 | and 38: "~x1 | ~x15" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 853 | and 39: "~x1 | ~x22" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 854 | and 40: "~x1 | ~x29" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 855 | and 41: "~x1 | ~x36" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 856 | and 42: "~x1 | ~x43" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 857 | and 43: "~x1 | ~x50" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 858 | and 44: "~x8 | ~x15" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 859 | and 45: "~x8 | ~x22" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 860 | and 46: "~x8 | ~x29" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 861 | and 47: "~x8 | ~x36" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 862 | and 48: "~x8 | ~x43" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 863 | and 49: "~x8 | ~x50" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 864 | and 50: "~x15 | ~x22" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 865 | and 51: "~x15 | ~x29" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 866 | and 52: "~x15 | ~x36" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 867 | and 53: "~x15 | ~x43" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 868 | and 54: "~x15 | ~x50" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 869 | and 55: "~x22 | ~x29" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 870 | and 56: "~x22 | ~x36" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 871 | and 57: "~x22 | ~x43" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 872 | and 58: "~x22 | ~x50" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 873 | and 59: "~x29 | ~x36" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 874 | and 60: "~x29 | ~x43" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 875 | and 61: "~x29 | ~x50" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 876 | and 62: "~x36 | ~x43" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 877 | and 63: "~x36 | ~x50" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 878 | and 64: "~x43 | ~x50" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 879 | and 65: "~x2 | ~x9" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 880 | and 66: "~x2 | ~x16" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 881 | and 67: "~x2 | ~x23" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 882 | and 68: "~x2 | ~x30" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 883 | and 69: "~x2 | ~x37" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 884 | and 70: "~x2 | ~x44" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 885 | and 71: "~x2 | ~x51" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 886 | and 72: "~x9 | ~x16" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 887 | and 73: "~x9 | ~x23" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 888 | and 74: "~x9 | ~x30" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 889 | and 75: "~x9 | ~x37" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 890 | and 76: "~x9 | ~x44" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 891 | and 77: "~x9 | ~x51" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 892 | and 78: "~x16 | ~x23" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 893 | and 79: "~x16 | ~x30" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 894 | and 80: "~x16 | ~x37" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 895 | and 81: "~x16 | ~x44" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 896 | and 82: "~x16 | ~x51" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 897 | and 83: "~x23 | ~x30" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 898 | and 84: "~x23 | ~x37" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 899 | and 85: "~x23 | ~x44" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 900 | and 86: "~x23 | ~x51" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 901 | and 87: "~x30 | ~x37" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 902 | and 88: "~x30 | ~x44" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 903 | and 89: "~x30 | ~x51" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 904 | and 90: "~x37 | ~x44" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 905 | and 91: "~x37 | ~x51" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 906 | and 92: "~x44 | ~x51" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 907 | and 93: "~x3 | ~x10" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 908 | and 94: "~x3 | ~x17" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 909 | and 95: "~x3 | ~x24" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 910 | and 96: "~x3 | ~x31" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 911 | and 97: "~x3 | ~x38" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 912 | and 98: "~x3 | ~x45" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 913 | and 99: "~x3 | ~x52" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 914 | and 100: "~x10 | ~x17" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 915 | and 101: "~x10 | ~x24" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 916 | and 102: "~x10 | ~x31" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 917 | and 103: "~x10 | ~x38" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 918 | and 104: "~x10 | ~x45" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 919 | and 105: "~x10 | ~x52" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 920 | and 106: "~x17 | ~x24" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 921 | and 107: "~x17 | ~x31" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 922 | and 108: "~x17 | ~x38" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 923 | and 109: "~x17 | ~x45" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 924 | and 110: "~x17 | ~x52" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 925 | and 111: "~x24 | ~x31" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 926 | and 112: "~x24 | ~x38" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 927 | and 113: "~x24 | ~x45" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 928 | and 114: "~x24 | ~x52" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 929 | and 115: "~x31 | ~x38" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 930 | and 116: "~x31 | ~x45" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 931 | and 117: "~x31 | ~x52" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 932 | and 118: "~x38 | ~x45" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 933 | and 119: "~x38 | ~x52" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 934 | and 120: "~x45 | ~x52" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 935 | and 121: "~x4 | ~x11" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 936 | and 122: "~x4 | ~x18" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 937 | and 123: "~x4 | ~x25" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 938 | and 124: "~x4 | ~x32" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 939 | and 125: "~x4 | ~x39" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 940 | and 126: "~x4 | ~x46" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 941 | and 127: "~x4 | ~x53" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 942 | and 128: "~x11 | ~x18" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 943 | and 129: "~x11 | ~x25" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 944 | and 130: "~x11 | ~x32" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 945 | and 131: "~x11 | ~x39" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 946 | and 132: "~x11 | ~x46" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 947 | and 133: "~x11 | ~x53" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 948 | and 134: "~x18 | ~x25" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 949 | and 135: "~x18 | ~x32" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 950 | and 136: "~x18 | ~x39" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 951 | and 137: "~x18 | ~x46" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 952 | and 138: "~x18 | ~x53" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 953 | and 139: "~x25 | ~x32" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 954 | and 140: "~x25 | ~x39" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 955 | and 141: "~x25 | ~x46" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 956 | and 142: "~x25 | ~x53" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 957 | and 143: "~x32 | ~x39" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 958 | and 144: "~x32 | ~x46" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 959 | and 145: "~x32 | ~x53" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 960 | and 146: "~x39 | ~x46" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 961 | and 147: "~x39 | ~x53" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 962 | and 148: "~x46 | ~x53" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 963 | and 149: "~x5 | ~x12" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 964 | and 150: "~x5 | ~x19" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 965 | and 151: "~x5 | ~x26" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 966 | and 152: "~x5 | ~x33" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 967 | and 153: "~x5 | ~x40" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 968 | and 154: "~x5 | ~x47" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 969 | and 155: "~x5 | ~x54" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 970 | and 156: "~x12 | ~x19" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 971 | and 157: "~x12 | ~x26" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 972 | and 158: "~x12 | ~x33" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 973 | and 159: "~x12 | ~x40" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 974 | and 160: "~x12 | ~x47" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 975 | and 161: "~x12 | ~x54" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 976 | and 162: "~x19 | ~x26" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 977 | and 163: "~x19 | ~x33" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 978 | and 164: "~x19 | ~x40" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 979 | and 165: "~x19 | ~x47" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 980 | and 166: "~x19 | ~x54" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 981 | and 167: "~x26 | ~x33" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 982 | and 168: "~x26 | ~x40" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 983 | and 169: "~x26 | ~x47" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 984 | and 170: "~x26 | ~x54" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 985 | and 171: "~x33 | ~x40" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 986 | and 172: "~x33 | ~x47" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 987 | and 173: "~x33 | ~x54" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 988 | and 174: "~x40 | ~x47" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 989 | and 175: "~x40 | ~x54" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 990 | and 176: "~x47 | ~x54" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 991 | and 177: "~x6 | ~x13" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 992 | and 178: "~x6 | ~x20" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 993 | and 179: "~x6 | ~x27" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 994 | and 180: "~x6 | ~x34" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 995 | and 181: "~x6 | ~x41" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 996 | and 182: "~x6 | ~x48" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 997 | and 183: "~x6 | ~x55" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 998 | and 184: "~x13 | ~x20" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 999 | and 185: "~x13 | ~x27" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1000 | and 186: "~x13 | ~x34" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1001 | and 187: "~x13 | ~x41" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1002 | and 188: "~x13 | ~x48" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1003 | and 189: "~x13 | ~x55" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1004 | and 190: "~x20 | ~x27" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1005 | and 191: "~x20 | ~x34" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1006 | and 192: "~x20 | ~x41" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1007 | and 193: "~x20 | ~x48" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1008 | and 194: "~x20 | ~x55" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1009 | and 195: "~x27 | ~x34" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1010 | and 196: "~x27 | ~x41" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1011 | and 197: "~x27 | ~x48" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1012 | and 198: "~x27 | ~x55" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1013 | and 199: "~x34 | ~x41" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1014 | and 200: "~x34 | ~x48" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1015 | and 201: "~x34 | ~x55" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1016 | and 202: "~x41 | ~x48" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1017 | and 203: "~x41 | ~x55" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1018 | and 204: "~x48 | ~x55" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1019 | shows "False" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1020 | using assms | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1021 | by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1022 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1023 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1024 | lemma "0 \<le> (yc::real) \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1025 | 0 \<le> yd \<and> 0 \<le> yb \<and> 0 \<le> ya \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1026 | 0 \<le> yf \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1027 | 0 \<le> xh \<and> 0 \<le> ye \<and> 0 \<le> yg \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1028 | 0 \<le> yw \<and> 0 \<le> xs \<and> 0 \<le> yu \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1029 | 0 \<le> aea \<and> 0 \<le> aee \<and> 0 \<le> aed \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1030 | 0 \<le> zy \<and> 0 \<le> xz \<and> 0 \<le> zw \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1031 | 0 \<le> zb \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1032 | 0 \<le> za \<and> 0 \<le> yy \<and> 0 \<le> yz \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1033 | 0 \<le> zp \<and> 0 \<le> zo \<and> 0 \<le> yq \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1034 | 0 \<le> adp \<and> 0 \<le> aeb \<and> 0 \<le> aec \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1035 | 0 \<le> acm \<and> 0 \<le> aco \<and> 0 \<le> acn \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1036 | 0 \<le> abl \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1037 | 0 \<le> zr \<and> 0 \<le> zq \<and> 0 \<le> abh \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1038 | 0 \<le> abq \<and> 0 \<le> zd \<and> 0 \<le> abo \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1039 | 0 \<le> acd \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1040 | 0 \<le> acc \<and> 0 \<le> xi \<and> 0 \<le> acb \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1041 | 0 \<le> acp \<and> 0 \<le> acr \<and> 0 \<le> acq \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1042 | 0 \<le> xw \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1043 | 0 \<le> xr \<and> 0 \<le> xv \<and> 0 \<le> xu \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1044 | 0 \<le> zc \<and> 0 \<le> acg \<and> 0 \<le> ach \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1045 | 0 \<le> zt \<and> 0 \<le> zs \<and> 0 \<le> xy \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1046 | 0 \<le> ady \<and> 0 \<le> adw \<and> 0 \<le> zg \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1047 | 0 \<le> abd \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1048 | 0 \<le> abc \<and> 0 \<le> yr \<and> 0 \<le> abb \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1049 | 0 \<le> adi \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1050 | 0 \<le> x \<and> 0 \<le> adh \<and> 0 \<le> xa \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1051 | 0 \<le> aak \<and> 0 \<le> aai \<and> 0 \<le> aad \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1052 | 0 \<le> aba \<and> 0 \<le> zh \<and> 0 \<le> aay \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1053 | 0 \<le> abg \<and> 0 \<le> ys \<and> 0 \<le> abe \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1054 | 0 \<le> abs1 \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1055 | 0 \<le> yt \<and> 0 \<le> abr \<and> 0 \<le> zu \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1056 | 0 \<le> abv \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1057 | 0 \<le> zn \<and> 0 \<le> abw \<and> 0 \<le> zm \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1058 | 0 \<le> adl \<and> 0 \<le> adn \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1059 | 0 \<le> acf \<and> 0 \<le> aca \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1060 | 0 \<le> ads \<and> 0 \<le> aaq \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1061 | 0 \<le> ada \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1062 | 0 \<le> aaf \<and> 0 \<le> aac \<and> 0 \<le> aag \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1063 | 0 \<le> aal \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1064 | 0 \<le> acu \<and> 0 \<le> acs \<and> 0 \<le> act \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1065 | 0 \<le> aas \<and> 0 \<le> xb \<and> 0 \<le> aat \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1066 | 0 \<le> zk \<and> 0 \<le> zj \<and> 0 \<le> zi \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1067 | 0 \<le> ack \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1068 | 0 \<le> acj \<and> 0 \<le> xc \<and> 0 \<le> aci \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1069 | 0 \<le> aav \<and> 0 \<le> aah \<and> 0 \<le> xd \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1070 | 0 \<le> abt \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1071 | 0 \<le> xo \<and> 0 \<le> abu \<and> 0 \<le> xn \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1072 | 0 \<le> adc \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1073 | 0 \<le> abz \<and> 0 \<le> adc \<and> 0 \<le> abz \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1074 | 0 \<le> xt \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1075 | 0 \<le> zz \<and> 0 \<le> aab \<and> 0 \<le> aaa \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1076 | 0 \<le> adq \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1077 | 0 \<le> xl \<and> 0 \<le> adr \<and> 0 \<le> adb \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1078 | 0 \<le> zf \<and> 0 \<le> yh \<and> 0 \<le> yi \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1079 | 0 \<le> aao \<and> 0 \<le> aam \<and> 0 \<le> xe \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1080 | 0 \<le> abk \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1081 | 0 \<le> aby \<and> 0 \<le> abj \<and> 0 \<le> abx \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1082 | 0 \<le> yp \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1083 | 0 \<le> yl \<and> 0 \<le> yj \<and> 0 \<le> ym \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1084 | 0 \<le> acw \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1085 | 0 \<le> adk \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1086 | 0 \<le> adg \<and> 0 \<le> adj \<and> 0 \<le> adf \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1087 | 0 \<le> adv \<and> 0 \<le> xf \<and> 0 \<le> adu \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1088 | yc + yd + yb + ya = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1089 | yf + xh + ye + yg = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1090 | yw + xs + yu = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1091 | aea + aee + aed = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1092 | zy + xz + zw = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1093 | zb + za + yy + yz = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1094 | zp + zo + yq = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1095 | adp + aeb + aec = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1096 | acm + aco + acn = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1097 | abl + abl = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1098 | zr + zq + abh = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1099 | abq + zd + abo = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1100 | acd + acc + xi + acb = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1101 | acp + acr + acq = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1102 | xw + xr + xv + xu = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1103 | zc + acg + ach = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1104 | zt + zs + xy = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1105 | ady + adw + zg = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1106 | abd + abc + yr + abb = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1107 | adi + x + adh + xa = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1108 | aak + aai + aad = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1109 | aba + zh + aay = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1110 | abg + ys + abe = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1111 | abs1 + yt + abr + zu = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1112 | abv + zn + abw + zm = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1113 | adl + adn = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1114 | acf + aca = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1115 | ads + aaq = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1116 | ada + ada = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1117 | aaf + aac + aag = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1118 | aal + acu + acs + act = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1119 | aas + xb + aat = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1120 | zk + zj + zi = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1121 | ack + acj + xc + aci = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1122 | aav + aah + xd = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1123 | abt + xo + abu + xn = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1124 | adc + abz + adc + abz = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1125 | xt + zz + aab + aaa = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1126 | adq + xl + adr + adb = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1127 | zf + yh + yi = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1128 | aao + aam + xe = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1129 | abk + aby + abj + abx = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1130 | yp + yp = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1131 | yl + yj + ym = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1132 | acw + acw + acw + acw = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1133 | adk + adg + adj + adf = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1134 | adv + xf + adu = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1135 | yd = 0 \<or> yb = 0 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1136 | xh = 0 \<or> ye = 0 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1137 | yy = 0 \<or> za = 0 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1138 | acc = 0 \<or> xi = 0 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1139 | xv = 0 \<or> xr = 0 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1140 | yr = 0 \<or> abc = 0 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1141 | zn = 0 \<or> abw = 0 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1142 | xo = 0 \<or> abu = 0 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1143 | xl = 0 \<or> adr = 0 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1144 | (yr + abd < abl \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1145 | yr + (abd + abb) < 1) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1146 | yr + abd = abl \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1147 | yr + (abd + abb) = 1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1148 | adb + adr < xn + abu \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1149 | adb + adr = xn + abu \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1150 | (abl < abt \<or> abl < abt + xo) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1151 | abl = abt \<and> abl = abt + xo \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1152 | yd + yc < abc + abd \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1153 | yd + yc = abc + abd \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1154 | aca < abb + yr \<or> aca = abb + yr \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1155 | acb + acc < xu + xv \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1156 | acb + acc = xu + xv \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1157 | (yq < xu + xr \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1158 | yq + zp < xu + (xr + xw)) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1159 | yq = xu + xr \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1160 | yq + zp = xu + (xr + xw) \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1161 | (zw < xw \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1162 | zw < xw + xv \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1163 | zw + zy < xw + (xv + xu)) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1164 | zw = xw \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1165 | zw = xw + xv \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1166 | zw + zy = xw + (xv + xu) \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1167 | xs + yw < zs + zt \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1168 | xs + yw = zs + zt \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1169 | aab + xt < ye + yf \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1170 | aab + xt = ye + yf \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1171 | (ya + yb < yg + ye \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1172 | ya + (yb + yc) < yg + (ye + yf)) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1173 | ya + yb = yg + ye \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1174 | ya + (yb + yc) = yg + (ye + yf) \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1175 | (xu + xv < acb + acc \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1176 | xu + (xv + xw) < acb + (acc + acd)) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1177 | xu + xv = acb + acc \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1178 | xu + (xv + xw) = acb + (acc + acd) \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1179 | (zs < xz + zy \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1180 | zs + xy < xz + (zy + zw)) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1181 | zs = xz + zy \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1182 | zs + xy = xz + (zy + zw) \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1183 | (zs + zt < xz + zy \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1184 | zs + (zt + xy) < xz + (zy + zw)) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1185 | zs + zt = xz + zy \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1186 | zs + (zt + xy) = xz + (zy + zw) \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1187 | yg + ye < ya + yb \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1188 | yg + ye = ya + yb \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1189 | (abd < yc \<or> abd + abc < yc + yd) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1190 | abd = yc \<and> abd + abc = yc + yd \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1191 | (ye + yf < adr + adq \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1192 | ye + (yf + yg) < adr + (adq + adb)) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1193 | ye + yf = adr + adq \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1194 | ye + (yf + yg) = adr + (adq + adb) \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1195 | yh + yi < ym + yj \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1196 | yh + yi = ym + yj \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1197 | (abq < yl \<or> abq + abo < yl + ym) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1198 | abq = yl \<and> abq + abo = yl + ym \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1199 | (yp < zp \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1200 | yp < zp + zo \<or> 1 < zp + (zo + yq)) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1201 | yp = zp \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1202 | yp = zp + zo \<and> 1 = zp + (zo + yq) \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1203 | (abb + yr < aca \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1204 | abb + (yr + abd) < aca + acf) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1205 | abb + yr = aca \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1206 | abb + (yr + abd) = aca + acf \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1207 | adw + zg < abe + ys \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1208 | adw + zg = abe + ys \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1209 | zd + abq < ys + abg \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1210 | zd + abq = ys + abg \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1211 | yt + abs1 < aby + abk \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1212 | yt + abs1 = aby + abk \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1213 | (yu < abx \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1214 | yu < abx + aby \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1215 | yu + yw < abx + (aby + abk)) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1216 | yu = abx \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1217 | yu = abx + aby \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1218 | yu + yw = abx + (aby + abk) \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1219 | aaf < adv \<or> aaf = adv \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1220 | abj + abk < yy + zb \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1221 | abj + abk = yy + zb \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1222 | (abb < yz \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1223 | abb + abc < yz + za \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1224 | abb + (abc + abd) < yz + (za + zb)) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1225 | abb = yz \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1226 | abb + abc = yz + za \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1227 | abb + (abc + abd) = yz + (za + zb) \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1228 | (acg + zc < zd + abq \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1229 | acg + (zc + ach) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1230 | < zd + (abq + abo)) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1231 | acg + zc = zd + abq \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1232 | acg + (zc + ach) = | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1233 | zd + (abq + abo) \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1234 | zf < acm \<or> zf = acm \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1235 | (zg + ady < acn + acm \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1236 | zg + (ady + adw) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1237 | < acn + (acm + aco)) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1238 | zg + ady = acn + acm \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1239 | zg + (ady + adw) = | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1240 | acn + (acm + aco) \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1241 | aay + zh < zi + zj \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1242 | aay + zh = zi + zj \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1243 | zy < zk \<or> zy = zk \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1244 | (adn < zm + zn \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1245 | adn + adl < zm + (zn + abv)) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1246 | adn = zm + zn \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1247 | adn + adl = zm + (zn + abv) \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1248 | zo + zp < zs + zt \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1249 | zo + zp = zs + zt \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1250 | zq + zr < zs + zt \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1251 | zq + zr = zs + zt \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1252 | (aai < adi \<or> aai < adi + adh) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1253 | aai = adi \<and> aai = adi + adh \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1254 | (abr < acj \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1255 | abr + (abs1 + zu) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1256 | < acj + (aci + ack)) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1257 | abr = acj \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1258 | abr + (abs1 + zu) = | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1259 | acj + (aci + ack) \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1260 | (abl < zw \<or> 1 < zw + zy) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1261 | abl = zw \<and> 1 = zw + zy \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1262 | (zz + aaa < act + acu \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1263 | zz + (aaa + aab) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1264 | < act + (acu + aal)) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1265 | zz + aaa = act + acu \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1266 | zz + (aaa + aab) = | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1267 | act + (acu + aal) \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1268 | (aam < aac \<or> aam + aao < aac + aaf) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1269 | aam = aac \<and> aam + aao = aac + aaf \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1270 | (aak < aaf \<or> aak + aad < aaf + aag) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1271 | aak = aaf \<and> aak + aad = aaf + aag \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1272 | (aah < aai \<or> aah + aav < aai + aak) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1273 | aah = aai \<and> aah + aav = aai + aak \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1274 | act + (acu + aal) < aam + aao \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1275 | act + (acu + aal) = aam + aao \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1276 | (ads < aat \<or> 1 < aat + aas) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1277 | ads = aat \<and> 1 = aat + aas \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1278 | (aba < aas \<or> aba + aay < aas + aat) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1279 | aba = aas \<and> aba + aay = aas + aat \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1280 | acm < aav \<or> acm = aav \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1281 | (ada < aay \<or> 1 < aay + aba) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1282 | ada = aay \<and> 1 = aay + aba \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1283 | abb + (abc + abd) < abe + abg \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1284 | abb + (abc + abd) = abe + abg \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1285 | (abh < abj \<or> abh < abj + abk) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1286 | abh = abj \<and> abh = abj + abk \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1287 | 1 < abo + abq \<or> 1 = abo + abq \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1288 | (acj < abr \<or> acj + aci < abr + abs1) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1289 | acj = abr \<and> acj + aci = abr + abs1 \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1290 | (abt < abv \<or> abt + abu < abv + abw) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1291 | abt = abv \<and> abt + abu = abv + abw \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1292 | (abx < adc \<or> abx + aby < adc + abz) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1293 | abx = adc \<and> abx + aby = adc + abz \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1294 | (acf < acd \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1295 | acf < acd + acc \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1296 | 1 < acd + (acc + acb)) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1297 | acf = acd \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1298 | acf = acd + acc \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1299 | 1 = acd + (acc + acb) \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1300 | acc + acd < acf \<or> acc + acd = acf \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1301 | (acg < acq \<or> acg + ach < acq + acr) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1302 | acg = acq \<and> acg + ach = acq + acr \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1303 | aci + (acj + ack) < acr + acp \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1304 | aci + (acj + ack) = acr + acp \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1305 | (acm < acp \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1306 | acm + acn < acp + acq \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1307 | acm + (acn + aco) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1308 | < acp + (acq + acr)) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1309 | acm = acp \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1310 | acm + acn = acp + acq \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1311 | acm + (acn + aco) = | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1312 | acp + (acq + acr) \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1313 | (acs + act < acw + acw \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1314 | acs + (act + acu) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1315 | < acw + (acw + acw)) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1316 | acs + act = acw + acw \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1317 | acs + (act + acu) = | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1318 | acw + (acw + acw) \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1319 | (ada < adb + adr \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1320 | 1 < adb + (adr + adq)) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1321 | ada = adb + adr \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1322 | 1 = adb + (adr + adq) \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1323 | (adc + adc < adf + adg \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1324 | adc + (adc + abz) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1325 | < adf + (adg + adk)) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1326 | adc + adc = adf + adg \<and> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1327 | adc + (adc + abz) = | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1328 | adf + (adg + adk) \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1329 | adh + adi < adj + adk \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1330 | adh + adi = adj + adk \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1331 | (adl < aec \<or> 1 < aec + adp) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1332 | adl = aec \<and> 1 = aec + adp \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1333 | (adq < ads \<or> adq + adr < ads) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1334 | adq = ads \<and> adq + adr = ads \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1335 | adu + adv < aed + aea \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1336 | adu + adv = aed + aea \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1337 | (adw < aee \<or> adw + ady < aee + aea) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1338 | adw = aee \<and> adw + ady = aee + aea \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1339 | (aeb < aed \<or> aeb + aec < aed + aee) \<or> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1340 | aeb = aed \<and> aeb + aec = aed + aee \<Longrightarrow> | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1341 | False" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1342 | by argo | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1343 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 1344 | end |