diff -r 46f59511b7bb -r d97fdabd9e2b src/Pure/Examples/First_Order_Logic.thy --- a/src/Pure/Examples/First_Order_Logic.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/Pure/Examples/First_Order_Logic.thy Fri Sep 20 19:51:08 2024 +0200 @@ -18,21 +18,21 @@ typedecl i typedecl o -judgment Trueprop :: "o \ prop" ("_" 5) +judgment Trueprop :: "o \ prop" (\_\ 5) subsection \Propositional logic\ -axiomatization false :: o ("\") +axiomatization false :: o (\\\) where falseE [elim]: "\ \ A" -axiomatization imp :: "o \ o \ o" (infixr "\" 25) +axiomatization imp :: "o \ o \ o" (infixr \\\ 25) where impI [intro]: "(A \ B) \ A \ B" and mp [dest]: "A \ B \ A \ B" -axiomatization conj :: "o \ o \ o" (infixr "\" 35) +axiomatization conj :: "o \ o \ o" (infixr \\\ 35) where conjI [intro]: "A \ B \ A \ B" and conjD1: "A \ B \ A" and conjD2: "A \ B \ B" @@ -48,20 +48,20 @@ qed -axiomatization disj :: "o \ o \ o" (infixr "\" 30) +axiomatization disj :: "o \ o \ o" (infixr \\\ 30) where disjE [elim]: "A \ B \ (A \ C) \ (B \ C) \ C" and disjI1 [intro]: "A \ A \ B" and disjI2 [intro]: "B \ A \ B" -definition true :: o ("\") +definition true :: o (\\\) where "\ \ \ \ \" theorem trueI [intro]: \ unfolding true_def .. -definition not :: "o \ o" ("\ _" [40] 40) +definition not :: "o \ o" (\\ _\ [40] 40) where "\ A \ A \ \" theorem notI [intro]: "(A \ \) \ \ A" @@ -76,7 +76,7 @@ qed -definition iff :: "o \ o \ o" (infixr "\" 25) +definition iff :: "o \ o \ o" (infixr \\\ 25) where "A \ B \ (A \ B) \ (B \ A)" theorem iffI [intro]: @@ -112,7 +112,7 @@ subsection \Equality\ -axiomatization equal :: "i \ i \ o" (infixl "=" 50) +axiomatization equal :: "i \ i \ o" (infixl \=\ 50) where refl [intro]: "x = x" and subst: "x = y \ P x \ P y" @@ -129,11 +129,11 @@ subsection \Quantifiers\ -axiomatization All :: "(i \ o) \ o" (binder "\" 10) +axiomatization All :: "(i \ o) \ o" (binder \\\ 10) where allI [intro]: "(\x. P x) \ \x. P x" and allD [dest]: "\x. P x \ P a" -axiomatization Ex :: "(i \ o) \ o" (binder "\" 10) +axiomatization Ex :: "(i \ o) \ o" (binder \\\ 10) where exI [intro]: "P a \ \x. P x" and exE [elim]: "\x. P x \ (\x. P x \ C) \ C"