| author | paulson | 
| Wed, 10 Jan 2001 11:00:17 +0100 | |
| changeset 10840 | 28a53b68a8c0 | 
| parent 10311 | 3b53ed2c846f | 
| child 11089 | 0f6f1cd500e5 | 
| permissions | -rw-r--r-- | 
| 6779 
2912aff958bd
Calculation.thy: Setup transitivity rules for calculational proofs.
 wenzelm parents: diff
changeset | 1 | (* Title: HOL/Calculation.thy | 
| 
2912aff958bd
Calculation.thy: Setup transitivity rules for calculational proofs.
 wenzelm parents: diff
changeset | 2 | ID: $Id$ | 
| 
2912aff958bd
Calculation.thy: Setup transitivity rules for calculational proofs.
 wenzelm parents: diff
changeset | 3 | Author: Markus Wenzel, TU Muenchen | 
| 9228 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 4 | License: GPL (GNU GENERAL PUBLIC LICENSE) | 
| 6779 
2912aff958bd
Calculation.thy: Setup transitivity rules for calculational proofs.
 wenzelm parents: diff
changeset | 5 | |
| 9228 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 6 | Setup transitivity rules for calculational proofs. | 
| 6779 
2912aff958bd
Calculation.thy: Setup transitivity rules for calculational proofs.
 wenzelm parents: diff
changeset | 7 | *) | 
| 
2912aff958bd
Calculation.thy: Setup transitivity rules for calculational proofs.
 wenzelm parents: diff
changeset | 8 | |
| 9035 | 9 | theory Calculation = IntArith: | 
| 6779 
2912aff958bd
Calculation.thy: Setup transitivity rules for calculational proofs.
 wenzelm parents: diff
changeset | 10 | |
| 9482 | 11 | lemma forw_subst: "a = b ==> P b ==> P a" | 
| 9035 | 12 | by (rule ssubst) | 
| 6873 | 13 | |
| 9482 | 14 | lemma back_subst: "P a ==> a = b ==> P b" | 
| 9035 | 15 | by (rule subst) | 
| 7381 | 16 | |
| 9482 | 17 | lemma set_rev_mp: "x:A ==> A <= B ==> x:B" | 
| 9035 | 18 | by (rule subsetD) | 
| 7657 | 19 | |
| 9482 | 20 | lemma set_mp: "A <= B ==> x:A ==> x:B" | 
| 9035 | 21 | by (rule subsetD) | 
| 7657 | 22 | |
| 9482 | 23 | lemma order_neq_le_trans: "a ~= b ==> (a::'a::order) <= b ==> a < b" | 
| 9228 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 24 | by (simp add: order_less_le) | 
| 7561 | 25 | |
| 9482 | 26 | lemma order_le_neq_trans: "(a::'a::order) <= b ==> a ~= b ==> a < b" | 
| 9228 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 27 | by (simp add: order_less_le) | 
| 7561 | 28 | |
| 9482 | 29 | lemma order_less_asym': "(a::'a::order) < b ==> b < a ==> P" | 
| 9035 | 30 | by (rule order_less_asym) | 
| 6873 | 31 | |
| 9482 | 32 | lemma ord_le_eq_trans: "a <= b ==> b = c ==> a <= c" | 
| 9228 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 33 | by (rule subst) | 
| 6862 | 34 | |
| 9482 | 35 | lemma ord_eq_le_trans: "a = b ==> b <= c ==> a <= c" | 
| 9228 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 36 | by (rule ssubst) | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 37 | |
| 9482 | 38 | lemma ord_less_eq_trans: "a < b ==> b = c ==> a < c" | 
| 9035 | 39 | by (rule subst) | 
| 6779 
2912aff958bd
Calculation.thy: Setup transitivity rules for calculational proofs.
 wenzelm parents: diff
changeset | 40 | |
| 9482 | 41 | lemma ord_eq_less_trans: "a = b ==> b < c ==> a < c" | 
| 9035 | 42 | by (rule ssubst) | 
| 6862 | 43 | |
| 9482 | 44 | lemma order_less_subst2: "(a::'a::order) < b ==> f b < (c::'c::order) ==> | 
| 9228 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 45 | (!!x y. x < y ==> f x < f y) ==> f a < c" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 46 | proof - | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 47 | assume r: "!!x y. x < y ==> f x < f y" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 48 | assume "a < b" hence "f a < f b" by (rule r) | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 49 | also assume "f b < c" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 50 | finally (order_less_trans) show ?thesis . | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 51 | qed | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 52 | |
| 9482 | 53 | lemma order_less_subst1: "(a::'a::order) < f b ==> (b::'b::order) < c ==> | 
| 9228 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 54 | (!!x y. x < y ==> f x < f y) ==> a < f c" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 55 | proof - | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 56 | assume r: "!!x y. x < y ==> f x < f y" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 57 | assume "a < f b" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 58 | also assume "b < c" hence "f b < f c" by (rule r) | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 59 | finally (order_less_trans) show ?thesis . | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 60 | qed | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 61 | |
| 9482 | 62 | lemma order_le_less_subst2: "(a::'a::order) <= b ==> f b < (c::'c::order) ==> | 
| 9228 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 63 | (!!x y. x <= y ==> f x <= f y) ==> f a < c" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 64 | proof - | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 65 | assume r: "!!x y. x <= y ==> f x <= f y" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 66 | assume "a <= b" hence "f a <= f b" by (rule r) | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 67 | also assume "f b < c" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 68 | finally (order_le_less_trans) show ?thesis . | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 69 | qed | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 70 | |
| 9482 | 71 | lemma order_le_less_subst1: "(a::'a::order) <= f b ==> (b::'b::order) < c ==> | 
| 9228 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 72 | (!!x y. x < y ==> f x < f y) ==> a < f c" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 73 | proof - | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 74 | assume r: "!!x y. x < y ==> f x < f y" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 75 | assume "a <= f b" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 76 | also assume "b < c" hence "f b < f c" by (rule r) | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 77 | finally (order_le_less_trans) show ?thesis . | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 78 | qed | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 79 | |
| 9482 | 80 | lemma order_less_le_subst2: "(a::'a::order) < b ==> f b <= (c::'c::order) ==> | 
| 9228 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 81 | (!!x y. x < y ==> f x < f y) ==> f a < c" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 82 | proof - | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 83 | assume r: "!!x y. x < y ==> f x < f y" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 84 | assume "a < b" hence "f a < f b" by (rule r) | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 85 | also assume "f b <= c" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 86 | finally (order_less_le_trans) show ?thesis . | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 87 | qed | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 88 | |
| 9482 | 89 | lemma order_less_le_subst1: "(a::'a::order) < f b ==> (b::'b::order) <= c ==> | 
| 9228 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 90 | (!!x y. x <= y ==> f x <= f y) ==> a < f c" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 91 | proof - | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 92 | assume r: "!!x y. x <= y ==> f x <= f y" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 93 | assume "a < f b" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 94 | also assume "b <= c" hence "f b <= f c" by (rule r) | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 95 | finally (order_less_le_trans) show ?thesis . | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 96 | qed | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 97 | |
| 9482 | 98 | lemma order_subst1: "(a::'a::order) <= f b ==> (b::'b::order) <= c ==> | 
| 9228 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 99 | (!!x y. x <= y ==> f x <= f y) ==> a <= f c" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 100 | proof - | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 101 | assume r: "!!x y. x <= y ==> f x <= f y" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 102 | assume "a <= f b" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 103 | also assume "b <= c" hence "f b <= f c" by (rule r) | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 104 | finally (order_trans) show ?thesis . | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 105 | qed | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 106 | |
| 9482 | 107 | lemma order_subst2: "(a::'a::order) <= b ==> f b <= (c::'c::order) ==> | 
| 9228 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 108 | (!!x y. x <= y ==> f x <= f y) ==> f a <= c" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 109 | proof - | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 110 | assume r: "!!x y. x <= y ==> f x <= f y" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 111 | assume "a <= b" hence "f a <= f b" by (rule r) | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 112 | also assume "f b <= c" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 113 | finally (order_trans) show ?thesis . | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 114 | qed | 
| 6779 
2912aff958bd
Calculation.thy: Setup transitivity rules for calculational proofs.
 wenzelm parents: diff
changeset | 115 | |
| 9482 | 116 | lemma ord_le_eq_subst: "a <= b ==> f b = c ==> | 
| 9228 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 117 | (!!x y. x <= y ==> f x <= f y) ==> f a <= c" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 118 | proof - | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 119 | assume r: "!!x y. x <= y ==> f x <= f y" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 120 | assume "a <= b" hence "f a <= f b" by (rule r) | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 121 | also assume "f b = c" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 122 | finally (ord_le_eq_trans) show ?thesis . | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 123 | qed | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 124 | |
| 9482 | 125 | lemma ord_eq_le_subst: "a = f b ==> b <= c ==> | 
| 9228 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 126 | (!!x y. x <= y ==> f x <= f y) ==> a <= f c" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 127 | proof - | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 128 | assume r: "!!x y. x <= y ==> f x <= f y" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 129 | assume "a = f b" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 130 | also assume "b <= c" hence "f b <= f c" by (rule r) | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 131 | finally (ord_eq_le_trans) show ?thesis . | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 132 | qed | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 133 | |
| 9482 | 134 | lemma ord_less_eq_subst: "a < b ==> f b = c ==> | 
| 9228 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 135 | (!!x y. x < y ==> f x < f y) ==> f a < c" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 136 | proof - | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 137 | assume r: "!!x y. x < y ==> f x < f y" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 138 | assume "a < b" hence "f a < f b" by (rule r) | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 139 | also assume "f b = c" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 140 | finally (ord_less_eq_trans) show ?thesis . | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 141 | qed | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 142 | |
| 9482 | 143 | lemma ord_eq_less_subst: "a = f b ==> b < c ==> | 
| 9228 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 144 | (!!x y. x < y ==> f x < f y) ==> a < f c" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 145 | proof - | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 146 | assume r: "!!x y. x < y ==> f x < f y" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 147 | assume "a = f b" | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 148 | also assume "b < c" hence "f b < f c" by (rule r) | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 149 | finally (ord_eq_less_trans) show ?thesis . | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 150 | qed | 
| 6779 
2912aff958bd
Calculation.thy: Setup transitivity rules for calculational proofs.
 wenzelm parents: diff
changeset | 151 | |
| 9228 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 152 | text {*
 | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 153 | Note that this list of rules is in reverse order of priorities. | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 154 | *} | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 155 | |
| 10311 | 156 | lemmas basic_trans_rules [trans] = | 
| 9228 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 157 | order_less_subst2 | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 158 | order_less_subst1 | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 159 | order_le_less_subst2 | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 160 | order_le_less_subst1 | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 161 | order_less_le_subst2 | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 162 | order_less_le_subst1 | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 163 | order_subst2 | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 164 | order_subst1 | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 165 | ord_le_eq_subst | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 166 | ord_eq_le_subst | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 167 | ord_less_eq_subst | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 168 | ord_eq_less_subst | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 169 | forw_subst | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 170 | back_subst | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 171 | dvd_trans | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 172 | rev_mp | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 173 | mp | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 174 | set_rev_mp | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 175 | set_mp | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 176 | order_neq_le_trans | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 177 | order_le_neq_trans | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 178 | order_less_asym' | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 179 | order_less_trans | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 180 | order_le_less_trans | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 181 | order_less_le_trans | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 182 | order_trans | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 183 | order_antisym | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 184 | ord_le_eq_trans | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 185 | ord_eq_le_trans | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 186 | ord_less_eq_trans | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 187 | ord_eq_less_trans | 
| 
672b03038110
added subst rules for ord(er), including monotonicity conditions;
 wenzelm parents: 
9142diff
changeset | 188 | trans | 
| 10130 | 189 | transitive | 
| 6945 | 190 | |
| 9035 | 191 | end |