diff -r 1d31e3a50373 -r 0318b43ee95c src/HOL/TLA/Intensional.thy --- a/src/HOL/TLA/Intensional.thy Fri Jun 26 11:07:04 2015 +0200 +++ b/src/HOL/TLA/Intensional.thy Fri Jun 26 11:44:22 2015 +0200 @@ -39,7 +39,7 @@ "" :: "var => lift" ("_") "_applC" :: "[lift, cargs] => lift" ("(1_/ _)" [1000, 1000] 999) "" :: "lift => lift" ("'(_')") - "_lambda" :: "[idts, 'a] => lift" ("(3%_./ _)" [0, 3] 3) + "_lambda" :: "[idts, 'a] => lift" ("(3\_./ _)" [0, 3] 3) "_constrain" :: "[lift, type] => lift" ("(_::_)" [4, 0] 3) "" :: "lift => liftargs" ("_") "_liftargs" :: "[lift, liftargs] => liftargs" ("_,/ _") @@ -155,29 +155,17 @@ "_liftMem" :: "[lift, lift] => lift" ("(_/ \ _)" [50, 51] 50) "_liftNotMem" :: "[lift, lift] => lift" ("(_/ \ _)" [50, 51] 50) -syntax (HTML output) - "_liftNeq" :: "[lift, lift] => lift" (infixl "\" 50) - "_liftNot" :: "lift => lift" ("\ _" [40] 40) - "_liftAnd" :: "[lift, lift] => lift" (infixr "\" 35) - "_liftOr" :: "[lift, lift] => lift" (infixr "\" 30) - "_RAll" :: "[idts, lift] => lift" ("(3\_./ _)" [0, 10] 10) - "_REx" :: "[idts, lift] => lift" ("(3\_./ _)" [0, 10] 10) - "_REx1" :: "[idts, lift] => lift" ("(3\!_./ _)" [0, 10] 10) - "_liftLeq" :: "[lift, lift] => lift" ("(_/ \ _)" [50, 51] 50) - "_liftMem" :: "[lift, lift] => lift" ("(_/ \ _)" [50, 51] 50) - "_liftNotMem" :: "[lift, lift] => lift" ("(_/ \ _)" [50, 51] 50) - defs - Valid_def: "|- A == ALL w. w |= A" + Valid_def: "|- A == \w. w |= A" unl_con: "LIFT #c w == c" unl_lift: "lift f x w == f (x w)" unl_lift2: "LIFT f w == f (x w) (y w)" unl_lift3: "LIFT f w == f (x w) (y w) (z w)" - unl_Rall: "w |= ALL x. A x == ALL x. (w |= A x)" - unl_Rex: "w |= EX x. A x == EX x. (w |= A x)" - unl_Rex1: "w |= EX! x. A x == EX! x. (w |= A x)" + unl_Rall: "w |= \x. A x == \x. (w |= A x)" + unl_Rex: "w |= \x. A x == \ x. (w |= A x)" + unl_Rex1: "w |= \!x. A x == \!x. (w |= A x)" subsection {* Lemmas and tactics for "intensional" logics. *} @@ -192,7 +180,7 @@ apply (erule spec) done -lemma intI [intro!]: "(!!w. w |= A) ==> |- A" +lemma intI [intro!]: "(\w. w |= A) ==> |- A" apply (unfold Valid_def) apply (rule allI) apply (erule meta_spec) @@ -207,21 +195,21 @@ lemma int_simps: "|- (x=x) = #True" - "|- (~#True) = #False" "|- (~#False) = #True" "|- (~~ P) = P" - "|- ((~P) = P) = #False" "|- (P = (~P)) = #False" - "|- (P ~= Q) = (P = (~Q))" + "|- (\#True) = #False" "|- (\#False) = #True" "|- (\\ P) = P" + "|- ((\P) = P) = #False" "|- (P = (\P)) = #False" + "|- (P \ Q) = (P = (\Q))" "|- (#True=P) = P" "|- (P=#True) = P" "|- (#True --> P) = P" "|- (#False --> P) = #True" "|- (P --> #True) = #True" "|- (P --> P) = #True" - "|- (P --> #False) = (~P)" "|- (P --> ~P) = (~P)" + "|- (P --> #False) = (\P)" "|- (P --> \P) = (\P)" "|- (P & #True) = P" "|- (#True & P) = P" "|- (P & #False) = #False" "|- (#False & P) = #False" - "|- (P & P) = P" "|- (P & ~P) = #False" "|- (~P & P) = #False" + "|- (P & P) = P" "|- (P & \P) = #False" "|- (\P & P) = #False" "|- (P | #True) = #True" "|- (#True | P) = #True" "|- (P | #False) = P" "|- (#False | P) = P" - "|- (P | P) = P" "|- (P | ~P) = #True" "|- (~P | P) = #True" - "|- (! x. P) = P" "|- (? x. P) = P" - "|- (~Q --> ~P) = (P --> Q)" + "|- (P | P) = P" "|- (P | \P) = #True" "|- (\P | P) = #True" + "|- (\x. P) = P" "|- (\x. P) = P" + "|- (\Q --> \P) = (P --> Q)" "|- (P|Q --> R) = ((P-->R)&(Q-->R))" apply (unfold Valid_def intensional_rews) apply blast+ @@ -296,10 +284,10 @@ attribute_setup int_use = {* Scan.succeed (Thm.rule_attribute (int_use o Context.proof_of)) *} -lemma Not_Rall: "|- (~(! x. F x)) = (? x. ~F x)" +lemma Not_Rall: "|- (\(\x. F x)) = (\x. \F x)" by (simp add: Valid_def) -lemma Not_Rex: "|- (~ (? x. F x)) = (! x. ~ F x)" +lemma Not_Rex: "|- (\ (\x. F x)) = (\x. \ F x)" by (simp add: Valid_def) end