| 16833 |      1 | 
 | 
|  |      2 | (* $Id$ *)
 | 
|  |      3 | 
 | 
| 15198 |      4 | local
 | 
|  |      5 | 
 | 
|  |      6 | val order_antisym_conv = thm "order_antisym_conv"
 | 
|  |      7 | val linorder_antisym_conv1 = thm "linorder_antisym_conv1"
 | 
|  |      8 | val linorder_antisym_conv2 = thm "linorder_antisym_conv2"
 | 
|  |      9 | val linorder_antisym_conv3 = thm "linorder_antisym_conv3"
 | 
|  |     10 | 
 | 
|  |     11 | fun prp t thm = (#prop(rep_thm thm) = t);
 | 
|  |     12 | 
 | 
|  |     13 | fun prove_antisym_le sg ss ((le as Const(_,T)) $ r $ s) =
 | 
|  |     14 |   let val prems = prems_of_ss ss;
 | 
| 19277 |     15 |       val less = Const("Orderings.less",T);
 | 
| 15198 |     16 |       val t = HOLogic.mk_Trueprop(le $ s $ r);
 | 
|  |     17 |   in case Library.find_first (prp t) prems of
 | 
| 15531 |     18 |        NONE =>
 | 
| 16833 |     19 |          let val t = HOLogic.mk_Trueprop(HOLogic.Not $ (less $ r $ s))
 | 
| 15198 |     20 |          in case Library.find_first (prp t) prems of
 | 
| 15531 |     21 |               NONE => NONE
 | 
|  |     22 |             | SOME thm => SOME(mk_meta_eq(thm RS linorder_antisym_conv1))
 | 
| 15198 |     23 |          end
 | 
| 15531 |     24 |      | SOME thm => SOME(mk_meta_eq(thm RS order_antisym_conv))
 | 
| 15198 |     25 |   end
 | 
| 15531 |     26 |   handle THM _ => NONE;
 | 
| 15198 |     27 | 
 | 
|  |     28 | fun prove_antisym_less sg ss (NotC $ ((less as Const(_,T)) $ r $ s)) =
 | 
|  |     29 |   let val prems = prems_of_ss ss;
 | 
| 19277 |     30 |       val le = Const("Orderings.less_eq",T);
 | 
| 15198 |     31 |       val t = HOLogic.mk_Trueprop(le $ r $ s);
 | 
|  |     32 |   in case Library.find_first (prp t) prems of
 | 
| 15531 |     33 |        NONE =>
 | 
| 15198 |     34 |          let val t = HOLogic.mk_Trueprop(NotC $ (less $ s $ r))
 | 
|  |     35 |          in case Library.find_first (prp t) prems of
 | 
| 15531 |     36 |               NONE => NONE
 | 
|  |     37 |             | SOME thm => SOME(mk_meta_eq(thm RS linorder_antisym_conv3))
 | 
| 15198 |     38 |          end
 | 
| 15531 |     39 |      | SOME thm => SOME(mk_meta_eq(thm RS linorder_antisym_conv2))
 | 
| 15198 |     40 |   end
 | 
| 15531 |     41 |   handle THM _ => NONE;
 | 
| 15198 |     42 | 
 | 
| 20836 |     43 | val antisym_le = Simplifier.simproc (the_context())
 | 
| 15198 |     44 |   "antisym le" ["(x::'a::order) <= y"] prove_antisym_le;
 | 
| 20836 |     45 | val antisym_less = Simplifier.simproc (the_context())
 | 
| 15198 |     46 |   "antisym less" ["~ (x::'a::linorder) < y"] prove_antisym_less;
 | 
|  |     47 | 
 | 
|  |     48 | in
 | 
|  |     49 | 
 | 
|  |     50 | val antisym_setup =
 | 
| 18708 |     51 |  (fn thy => (Simplifier.change_simpset_of thy
 | 
|  |     52 |   (fn ss => ss addsimprocs [antisym_le, antisym_less]); thy));
 | 
| 15198 |     53 | 
 | 
| 16833 |     54 | end
 |