# HG changeset patch # User wenzelm # Date 1412706857 -7200 # Node ID dbe216a75a4b4ccd4c99fcddf1d287eb49668250 # Parent d49f3181030e795e77194ee2460898a901a533f3 more symbols; diff -r d49f3181030e -r dbe216a75a4b src/Pure/Pure.thy --- a/src/Pure/Pure.thy Tue Oct 07 20:27:31 2014 +0200 +++ b/src/Pure/Pure.thy Tue Oct 07 20:34:17 2014 +0200 @@ -234,37 +234,37 @@ subsection \Meta-level connectives in assumptions\ lemma meta_mp: - assumes "PROP P ==> PROP Q" and "PROP P" + assumes "PROP P \ PROP Q" and "PROP P" shows "PROP Q" - by (rule \PROP P ==> PROP Q\ [OF \PROP P\]) + by (rule \PROP P \ PROP Q\ [OF \PROP P\]) lemmas meta_impE = meta_mp [elim_format] lemma meta_spec: - assumes "!!x. PROP P x" + assumes "\x. PROP P x" shows "PROP P x" - by (rule \!!x. PROP P x\) + by (rule \\x. PROP P x\) lemmas meta_allE = meta_spec [elim_format] lemma swap_params: - "(!!x y. PROP P x y) == (!!y x. PROP P x y)" .. + "(\x y. PROP P x y) \ (\y x. PROP P x y)" .. subsection \Meta-level conjunction\ lemma all_conjunction: - "(!!x. PROP A x &&& PROP B x) == ((!!x. PROP A x) &&& (!!x. PROP B x))" + "(\x. PROP A x &&& PROP B x) \ ((\x. PROP A x) &&& (\x. PROP B x))" proof - assume conj: "!!x. PROP A x &&& PROP B x" - show "(!!x. PROP A x) &&& (!!x. PROP B x)" + assume conj: "\x. PROP A x &&& PROP B x" + show "(\x. PROP A x) &&& (\x. PROP B x)" proof - fix x from conj show "PROP A x" by (rule conjunctionD1) from conj show "PROP B x" by (rule conjunctionD2) qed next - assume conj: "(!!x. PROP A x) &&& (!!x. PROP B x)" + assume conj: "(\x. PROP A x) &&& (\x. PROP B x)" fix x show "PROP A x &&& PROP B x" proof - @@ -274,17 +274,17 @@ qed lemma imp_conjunction: - "(PROP A ==> PROP B &&& PROP C) == ((PROP A ==> PROP B) &&& (PROP A ==> PROP C))" + "(PROP A \ PROP B &&& PROP C) \ ((PROP A \ PROP B) &&& (PROP A \ PROP C))" proof - assume conj: "PROP A ==> PROP B &&& PROP C" - show "(PROP A ==> PROP B) &&& (PROP A ==> PROP C)" + assume conj: "PROP A \ PROP B &&& PROP C" + show "(PROP A \ PROP B) &&& (PROP A \ PROP C)" proof - assume "PROP A" from conj [OF \PROP A\] show "PROP B" by (rule conjunctionD1) from conj [OF \PROP A\] show "PROP C" by (rule conjunctionD2) qed next - assume conj: "(PROP A ==> PROP B) &&& (PROP A ==> PROP C)" + assume conj: "(PROP A \ PROP B) &&& (PROP A \ PROP C)" assume "PROP A" show "PROP B &&& PROP C" proof - @@ -294,16 +294,16 @@ qed lemma conjunction_imp: - "(PROP A &&& PROP B ==> PROP C) == (PROP A ==> PROP B ==> PROP C)" + "(PROP A &&& PROP B \ PROP C) \ (PROP A \ PROP B \ PROP C)" proof - assume r: "PROP A &&& PROP B ==> PROP C" + assume r: "PROP A &&& PROP B \ PROP C" assume ab: "PROP A" "PROP B" show "PROP C" proof (rule r) from ab show "PROP A &&& PROP B" . qed next - assume r: "PROP A ==> PROP B ==> PROP C" + assume r: "PROP A \ PROP B \ PROP C" assume conj: "PROP A &&& PROP B" show "PROP C" proof (rule r)