337 let |
337 let |
338 fun excluded t = |
338 fun excluded t = |
339 (* exclude numeric types: linear arithmetic subsumes transitivity *) |
339 (* exclude numeric types: linear arithmetic subsumes transitivity *) |
340 let val T = type_of t |
340 let val T = type_of t |
341 in |
341 in |
342 T = HOLogic.natT orelse T = HOLogic.intT orelse T = HOLogic.realT |
342 T = HOLogic.natT orelse T = HOLogic.intT orelse T = HOLogic.realT |
343 end; |
343 end; |
344 fun rel (bin_op $ t1 $ t2) = |
344 fun rel (bin_op $ t1 $ t2) = |
345 if excluded t1 then NONE |
345 if excluded t1 then NONE |
346 else if Pattern.matches thy (eq, bin_op) then SOME (t1, "=", t2) |
346 else if Pattern.matches thy (eq, bin_op) then SOME (t1, "=", t2) |
347 else if Pattern.matches thy (le, bin_op) then SOME (t1, "<=", t2) |
347 else if Pattern.matches thy (le, bin_op) then SOME (t1, "<=", t2) |
348 else if Pattern.matches thy (less, bin_op) then SOME (t1, "<", t2) |
348 else if Pattern.matches thy (less, bin_op) then SOME (t1, "<", t2) |
349 else NONE |
349 else NONE |
350 | rel _ = NONE; |
350 | rel _ = NONE; |
351 fun dec (Const (@{const_name Not}, _) $ t) = (case rel t |
351 fun dec (Const (@{const_name Not}, _) $ t) = (case rel t |
352 of NONE => NONE |
352 of NONE => NONE |
353 | SOME (t1, rel, t2) => SOME (t1, "~" ^ rel, t2)) |
353 | SOME (t1, rel, t2) => SOME (t1, "~" ^ rel, t2)) |
354 | dec x = rel x; |
354 | dec x = rel x; |
355 in dec t end |
355 in dec t end |
356 | decomp thy _ = NONE; |
356 | decomp thy _ = NONE; |
357 in |
357 in |
358 case s of |
358 case s of |