diff -r af95a414136a -r 6abda9b6b9c1 src/HOL/SMT2.thy --- a/src/HOL/SMT2.thy Tue Jun 03 16:02:41 2014 +0200 +++ b/src/HOL/SMT2.thy Tue Jun 03 16:02:42 2014 +0200 @@ -298,20 +298,20 @@ if_True if_False not_not lemma [z3_new_rule]: - "(P \ Q) = (\(\P \ \Q))" - "(P \ Q) = (\(\Q \ \P))" - "(\P \ Q) = (\(P \ \Q))" - "(\P \ Q) = (\(\Q \ P))" - "(P \ \Q) = (\(\P \ Q))" - "(P \ \Q) = (\(Q \ \P))" - "(\P \ \Q) = (\(P \ Q))" - "(\P \ \Q) = (\(Q \ P))" + "(P \ Q) = (\ (\ P \ \ Q))" + "(P \ Q) = (\ (\ Q \ \ P))" + "(\ P \ Q) = (\ (P \ \ Q))" + "(\ P \ Q) = (\ (\ Q \ P))" + "(P \ \ Q) = (\ (\ P \ Q))" + "(P \ \ Q) = (\ (Q \ \ P))" + "(\ P \ \ Q) = (\ (P \ Q))" + "(\ P \ \ Q) = (\ (Q \ P))" by auto lemma [z3_new_rule]: - "(P \ Q) = (Q \ \P)" - "(\P \ Q) = (P \ Q)" - "(\P \ Q) = (Q \ P)" + "(P \ Q) = (Q \ \ P)" + "(\ P \ Q) = (P \ Q)" + "(\ P \ Q) = (Q \ P)" "(True \ P) = P" "(P \ True) = True" "(False \ P) = True" @@ -319,41 +319,41 @@ by auto lemma [z3_new_rule]: - "((P = Q) \ R) = (R | (Q = (\P)))" + "((P = Q) \ R) = (R | (Q = (\ P)))" by auto lemma [z3_new_rule]: - "(\True) = False" - "(\False) = True" + "(\ True) = False" + "(\ False) = True" "(x = x) = True" "(P = True) = P" "(True = P) = P" - "(P = False) = (\P)" - "(False = P) = (\P)" - "((\P) = P) = False" - "(P = (\P)) = False" - "((\P) = (\Q)) = (P = Q)" - "\(P = (\Q)) = (P = Q)" - "\((\P) = Q) = (P = Q)" - "(P \ Q) = (Q = (\P))" - "(P = Q) = ((\P \ Q) \ (P \ \Q))" - "(P \ Q) = ((\P \ \Q) \ (P \ Q))" + "(P = False) = (\ P)" + "(False = P) = (\ P)" + "((\ P) = P) = False" + "(P = (\ P)) = False" + "((\ P) = (\ Q)) = (P = Q)" + "\ (P = (\ Q)) = (P = Q)" + "\ ((\ P) = Q) = (P = Q)" + "(P \ Q) = (Q = (\ P))" + "(P = Q) = ((\ P \ Q) \ (P \ \ Q))" + "(P \ Q) = ((\ P \ \ Q) \ (P \ Q))" by auto lemma [z3_new_rule]: - "(if P then P else \P) = True" - "(if \P then \P else P) = True" + "(if P then P else \ P) = True" + "(if \ P then \ P else P) = True" "(if P then True else False) = P" - "(if P then False else True) = (\P)" - "(if P then Q else True) = ((\P) \ Q)" - "(if P then Q else True) = (Q \ (\P))" - "(if P then Q else \Q) = (P = Q)" - "(if P then Q else \Q) = (Q = P)" - "(if P then \Q else Q) = (P = (\Q))" - "(if P then \Q else Q) = ((\Q) = P)" - "(if \P then x else y) = (if P then y else x)" - "(if P then (if Q then x else y) else x) = (if P \ (\Q) then y else x)" - "(if P then (if Q then x else y) else x) = (if (\Q) \ P then y else x)" + "(if P then False else True) = (\ P)" + "(if P then Q else True) = ((\ P) \ Q)" + "(if P then Q else True) = (Q \ (\ P))" + "(if P then Q else \ Q) = (P = Q)" + "(if P then Q else \ Q) = (Q = P)" + "(if P then \ Q else Q) = (P = (\ Q))" + "(if P then \ Q else Q) = ((\ Q) = P)" + "(if \ P then x else y) = (if P then y else x)" + "(if P then (if Q then x else y) else x) = (if P \ (\ Q) then y else x)" + "(if P then (if Q then x else y) else x) = (if (\ Q) \ P then y else x)" "(if P then (if Q then x else y) else y) = (if P \ Q then x else y)" "(if P then (if Q then x else y) else y) = (if Q \ P then x else y)" "(if P then x else if P then y else z) = (if P then x else z)" @@ -375,31 +375,31 @@ lemma [z3_new_rule]: (* for def-axiom *) "P = Q \ P \ Q" - "P = Q \ \P \ \Q" - "(\P) = Q \ \P \ Q" - "(\P) = Q \ P \ \Q" - "P = (\Q) \ \P \ Q" - "P = (\Q) \ P \ \Q" - "P \ Q \ P \ \Q" - "P \ Q \ \P \ Q" - "P \ (\Q) \ P \ Q" - "(\P) \ Q \ P \ Q" - "P \ Q \ P \ (\Q)" - "P \ Q \ (\P) \ Q" - "P \ \Q \ P \ Q" - "\P \ Q \ P \ Q" + "P = Q \ \ P \ \ Q" + "(\ P) = Q \ \ P \ Q" + "(\ P) = Q \ P \ \ Q" + "P = (\ Q) \ \ P \ Q" + "P = (\ Q) \ P \ \ Q" + "P \ Q \ P \ \ Q" + "P \ Q \ \ P \ Q" + "P \ (\ Q) \ P \ Q" + "(\ P) \ Q \ P \ Q" + "P \ Q \ P \ (\ Q)" + "P \ Q \ (\ P) \ Q" + "P \ \ Q \ P \ Q" + "\ P \ Q \ P \ Q" "P \ y = (if P then x else y)" "P \ (if P then x else y) = y" - "\P \ x = (if P then x else y)" - "\P \ (if P then x else y) = x" - "P \ R \ \(if P then Q else R)" - "\P \ Q \ \(if P then Q else R)" - "\(if P then Q else R) \ \P \ Q" - "\(if P then Q else R) \ P \ R" - "(if P then Q else R) \ \P \ \Q" - "(if P then Q else R) \ P \ \R" - "(if P then \Q else R) \ \P \ Q" - "(if P then Q else \R) \ P \ R" + "\ P \ x = (if P then x else y)" + "\ P \ (if P then x else y) = x" + "P \ R \ \ (if P then Q else R)" + "\ P \ Q \ \ (if P then Q else R)" + "\ (if P then Q else R) \ \ P \ Q" + "\ (if P then Q else R) \ P \ R" + "(if P then Q else R) \ \ P \ \ Q" + "(if P then Q else R) \ P \ \ R" + "(if P then \ Q else R) \ \ P \ Q" + "(if P then Q else \ R) \ P \ R" by auto hide_type (open) pattern