# HG changeset patch # User wenzelm # Date 1437256990 -7200 # Node ID 36d9f215c9820941c571ee0f16e7e2657864c6e0 # Parent d8d85a8172b504fc96e84a8899b33382c3bdb1b9 more symbols; diff -r d8d85a8172b5 -r 36d9f215c982 src/HOL/HOL.thy --- a/src/HOL/HOL.thy Sat Jul 18 22:58:50 2015 +0200 +++ b/src/HOL/HOL.thy Sun Jul 19 00:03:10 2015 +0200 @@ -70,24 +70,24 @@ typedecl bool judgment - Trueprop :: "bool => prop" ("(_)" 5) + Trueprop :: "bool \ prop" ("(_)" 5) axiomatization - implies :: "[bool, bool] => bool" (infixr "-->" 25) and - eq :: "['a, 'a] => bool" (infixl "=" 50) and - The :: "('a => bool) => 'a" + implies :: "[bool, bool] \ bool" (infixr "-->" 25) and + eq :: "['a, 'a] \ bool" (infixl "=" 50) and + The :: "('a \ bool) \ 'a" consts True :: bool False :: bool - Not :: "bool => bool" ("~ _" [40] 40) + Not :: "bool \ bool" ("~ _" [40] 40) - conj :: "[bool, bool] => bool" (infixr "&" 35) - disj :: "[bool, bool] => bool" (infixr "|" 30) + conj :: "[bool, bool] \ bool" (infixr "&" 35) + disj :: "[bool, bool] \ bool" (infixr "|" 30) - All :: "('a => bool) => bool" (binder "ALL " 10) - Ex :: "('a => bool) => bool" (binder "EX " 10) - Ex1 :: "('a => bool) => bool" (binder "EX! " 10) + All :: "('a \ bool) \ bool" (binder "ALL " 10) + Ex :: "('a \ bool) \ bool" (binder "EX " 10) + Ex1 :: "('a \ bool) \ bool" (binder "EX! " 10) subsubsection \Additional concrete syntax\ @@ -96,8 +96,8 @@ eq (infix "=" 50) abbreviation - not_equal :: "['a, 'a] => bool" (infixl "~=" 50) where - "x ~= y == ~ (x = y)" + not_equal :: "['a, 'a] \ bool" (infixl "~=" 50) where + "x ~= y \ ~ (x = y)" notation (output) not_equal (infix "~=" 50) @@ -119,14 +119,14 @@ not_equal (infix "\" 50) abbreviation (iff) - iff :: "[bool, bool] => bool" (infixr "<->" 25) where - "A <-> B == A = B" + iff :: "[bool, bool] \ bool" (infixr "<->" 25) where + "A <-> B \ A = B" notation (xsymbols) iff (infixr "\" 25) -syntax "_The" :: "[pttrn, bool] => 'a" ("(3THE _./ _)" [0, 10] 10) -translations "THE x. P" == "CONST The (%x. P)" +syntax "_The" :: "[pttrn, bool] \ 'a" ("(3THE _./ _)" [0, 10] 10) +translations "THE x. P" \ "CONST The (\x. P)" print_translation \ [(@{const_syntax The}, fn _ => fn [Abs abs] => let val (x, t) = Syntax_Trans.atomic_abs_tr' abs @@ -135,19 +135,19 @@ nonterminal letbinds and letbind syntax - "_bind" :: "[pttrn, 'a] => letbind" ("(2_ =/ _)" 10) - "" :: "letbind => letbinds" ("_") - "_binds" :: "[letbind, letbinds] => letbinds" ("_;/ _") - "_Let" :: "[letbinds, 'a] => 'a" ("(let (_)/ in (_))" [0, 10] 10) + "_bind" :: "[pttrn, 'a] \ letbind" ("(2_ =/ _)" 10) + "" :: "letbind \ letbinds" ("_") + "_binds" :: "[letbind, letbinds] \ letbinds" ("_;/ _") + "_Let" :: "[letbinds, 'a] \ 'a" ("(let (_)/ in (_))" [0, 10] 10) nonterminal case_syn and cases_syn syntax - "_case_syntax" :: "['a, cases_syn] => 'b" ("(case _ of/ _)" 10) - "_case1" :: "['a, 'b] => case_syn" ("(2_ =>/ _)" 10) - "" :: "case_syn => cases_syn" ("_") - "_case2" :: "[case_syn, cases_syn] => cases_syn" ("_/ | _") + "_case_syntax" :: "['a, cases_syn] \ 'b" ("(case _ of/ _)" 10) + "_case1" :: "['a, 'b] \ case_syn" ("(2_ =>/ _)" 10) + "" :: "case_syn \ cases_syn" ("_") + "_case2" :: "[case_syn, cases_syn] \ cases_syn" ("_/ | _") syntax (xsymbols) - "_case1" :: "['a, 'b] => case_syn" ("(2_ \/ _)" 10) + "_case1" :: "['a, 'b] \ case_syn" ("(2_ \/ _)" 10) notation (xsymbols) All (binder "\" 10) and @@ -170,7 +170,7 @@ axiomatization where refl: "t = (t::'a)" and subst: "s = t \ P s \ P t" and - ext: "(!!x::'a. (f x ::'b) = g x) ==> (%x. f x) = (%x. g x)" + ext: "(\x::'a. (f x ::'b) = g x) \ (\x. f x) = (\x. g x)" -- \Extensionality is built into the meta-logic, and this rule expresses a related property. It is an eta-expanded version of the traditional rule, and similar to the ABS rule of HOL\ and @@ -178,31 +178,31 @@ the_eq_trivial: "(THE x. x = a) = (a::'a)" axiomatization where - impI: "(P ==> Q) ==> P-->Q" and - mp: "[| P-->Q; P |] ==> Q" and + impI: "(P \ Q) \ P \ Q" and + mp: "\P \ Q; P\ \ Q" and - iff: "(P-->Q) --> (Q-->P) --> (P=Q)" and - True_or_False: "(P=True) | (P=False)" + iff: "(P \ Q) \ (Q \ P) \ (P = Q)" and + True_or_False: "(P = True) \ (P = False)" defs - True_def: "True == ((%x::bool. x) = (%x. x))" - All_def: "All(P) == (P = (%x. True))" - Ex_def: "Ex(P) == !Q. (!x. P x --> Q) --> Q" - False_def: "False == (!P. P)" - not_def: "~ P == P-->False" - and_def: "P & Q == !R. (P-->Q-->R) --> R" - or_def: "P | Q == !R. (P-->R) --> (Q-->R) --> R" - Ex1_def: "Ex1(P) == ? x. P(x) & (! y. P(y) --> y=x)" + True_def: "True \ ((\x::bool. x) = (\x. x))" + All_def: "All P \ (P = (\x. True))" + Ex_def: "Ex P \ \Q. (\x. P x \ Q) \ Q" + False_def: "False \ (\P. P)" + not_def: "\ P \ P \ False" + and_def: "P \ Q \ \R. (P \ Q \ R) \ R" + or_def: "P \ Q \ \R. (P \ R) \ (Q \ R) \ R" + Ex1_def: "Ex1 P \ \x. P x \ (\y. P y \ y = x)" definition If :: "bool \ 'a \ 'a \ 'a" ("(if (_)/ then (_)/ else (_))" [0, 0, 10] 10) - where "If P x y \ (THE z::'a. (P=True --> z=x) & (P=False --> z=y))" + where "If P x y \ (THE z::'a. (P = True \ z = x) \ (P = False \ z = y))" definition Let :: "'a \ ('a \ 'b) \ 'b" where "Let s f \ f s" translations - "_Let (_binds b bs) e" == "_Let b (_Let bs e)" - "let x = a in e" == "CONST Let a (%x. e)" + "_Let (_binds b bs) e" \ "_Let b (_Let bs e)" + "let x = a in e" \ "CONST Let a (\x. e)" axiomatization undefined :: 'a @@ -213,20 +213,20 @@ subsubsection \Equality\ -lemma sym: "s = t ==> t = s" +lemma sym: "s = t \ t = s" by (erule subst) (rule refl) -lemma ssubst: "t = s ==> P s ==> P t" +lemma ssubst: "t = s \ P s \ P t" by (drule sym) (erule subst) -lemma trans: "[| r=s; s=t |] ==> r=t" +lemma trans: "\r = s; s = t\ \ r = t" by (erule subst) -lemma trans_sym [Pure.elim?]: "r = s ==> t = s ==> r = t" +lemma trans_sym [Pure.elim?]: "r = s \ t = s \ r = t" by (rule trans [OF _ sym]) lemma meta_eq_to_obj_eq: - assumes meq: "A == B" + assumes meq: "A \ B" shows "A = B" by (unfold meq) (rule refl) @@ -234,7 +234,7 @@ (* a = b | | c = d *) -lemma box_equals: "[| a=b; a=c; b=d |] ==> c=d" +lemma box_equals: "\a = b; a = c; b = d\ \ c = d" apply (rule trans) apply (rule trans) apply (rule sym) @@ -243,33 +243,33 @@ text \For calculational reasoning:\ -lemma forw_subst: "a = b ==> P b ==> P a" +lemma forw_subst: "a = b \ P b \ P a" by (rule ssubst) -lemma back_subst: "P a ==> a = b ==> P b" +lemma back_subst: "P a \ a = b \ P b" by (rule subst) subsubsection \Congruence rules for application\ text \Similar to @{text AP_THM} in Gordon's HOL.\ -lemma fun_cong: "(f::'a=>'b) = g ==> f(x)=g(x)" +lemma fun_cong: "(f :: 'a \ 'b) = g \ f x = g x" apply (erule subst) apply (rule refl) done text \Similar to @{text AP_TERM} in Gordon's HOL and FOL's @{text subst_context}.\ -lemma arg_cong: "x=y ==> f(x)=f(y)" +lemma arg_cong: "x = y \ f x = f y" apply (erule subst) apply (rule refl) done -lemma arg_cong2: "\ a = b; c = d \ \ f a c = f b d" +lemma arg_cong2: "\a = b; c = d\ \ f a c = f b d" apply (erule ssubst)+ apply (rule refl) done -lemma cong: "[| f = g; (x::'a) = y |] ==> f x = g y" +lemma cong: "\f = g; (x::'a) = y\ \ f x = g y" apply (erule subst)+ apply (rule refl) done @@ -279,13 +279,13 @@ subsubsection \Equality of booleans -- iff\ -lemma iffI: assumes "P ==> Q" and "Q ==> P" shows "P=Q" +lemma iffI: assumes "P \ Q" and "Q \ P" shows "P = Q" by (iprover intro: iff [THEN mp, THEN mp] impI assms) -lemma iffD2: "[| P=Q; Q |] ==> P" +lemma iffD2: "\P = Q; Q\ \ P" by (erule ssubst) -lemma rev_iffD2: "[| Q; P=Q |] ==> P" +lemma rev_iffD2: "\Q; P = Q\ \ P" by (erule iffD2) lemma iffD1: "Q = P \ Q \ P" @@ -295,8 +295,8 @@ by (drule sym) (rule rev_iffD2) lemma iffE: - assumes major: "P=Q" - and minor: "[| P --> Q; Q --> P |] ==> R" + assumes major: "P = Q" + and minor: "\P \ Q; Q \ P\ \ R" shows R by (iprover intro: minor impI major [THEN iffD2] major [THEN iffD1]) @@ -306,33 +306,33 @@ lemma TrueI: "True" unfolding True_def by (rule refl) -lemma eqTrueI: "P ==> P = True" +lemma eqTrueI: "P \ P = True" by (iprover intro: iffI TrueI) -lemma eqTrueE: "P = True ==> P" +lemma eqTrueE: "P = True \ P" by (erule iffD2) (rule TrueI) subsubsection \Universal quantifier\ -lemma allI: assumes "!!x::'a. P(x)" shows "ALL x. P(x)" +lemma allI: assumes "\x::'a. P x" shows "\x. P x" unfolding All_def by (iprover intro: ext eqTrueI assms) -lemma spec: "ALL x::'a. P(x) ==> P(x)" +lemma spec: "\x::'a. P x \ P x" apply (unfold All_def) apply (rule eqTrueE) apply (erule fun_cong) done lemma allE: - assumes major: "ALL x. P(x)" - and minor: "P(x) ==> R" + assumes major: "\x. P x" + and minor: "P x \ R" shows R by (iprover intro: minor major [THEN spec]) lemma all_dupE: - assumes major: "ALL x. P(x)" - and minor: "[| P(x); ALL x. P(x) |] ==> R" + assumes major: "\x. P x" + and minor: "\P x; \x. P x\ \ R" shows R by (iprover intro: minor major major [THEN spec]) @@ -344,36 +344,36 @@ logic before quantifiers! \ -lemma FalseE: "False ==> P" +lemma FalseE: "False \ P" apply (unfold False_def) apply (erule spec) done -lemma False_neq_True: "False = True ==> P" +lemma False_neq_True: "False = True \ P" by (erule eqTrueE [THEN FalseE]) subsubsection \Negation\ lemma notI: - assumes "P ==> False" - shows "~P" + assumes "P \ False" + shows "\ P" apply (unfold not_def) apply (iprover intro: impI assms) done -lemma False_not_True: "False ~= True" +lemma False_not_True: "False \ True" apply (rule notI) apply (erule False_neq_True) done -lemma True_not_False: "True ~= False" +lemma True_not_False: "True \ False" apply (rule notI) apply (drule sym) apply (erule False_neq_True) done -lemma notE: "[| ~P; P |] ==> R" +lemma notE: "\\ P; P\ \ R" apply (unfold not_def) apply (erule mp [THEN FalseE]) apply assumption @@ -386,44 +386,44 @@ subsubsection \Implication\ lemma impE: - assumes "P-->Q" "P" "Q ==> R" - shows "R" + assumes "P \ Q" P "Q \ R" + shows R by (iprover intro: assms mp) -(* Reduces Q to P-->Q, allowing substitution in P. *) -lemma rev_mp: "[| P; P --> Q |] ==> Q" +(* Reduces Q to P \ Q, allowing substitution in P. *) +lemma rev_mp: "\P; P \ Q\ \ Q" by (iprover intro: mp) lemma contrapos_nn: - assumes major: "~Q" - and minor: "P==>Q" - shows "~P" + assumes major: "\ Q" + and minor: "P \ Q" + shows "\ P" by (iprover intro: notI minor major [THEN notE]) (*not used at all, but we already have the other 3 combinations *) lemma contrapos_pn: assumes major: "Q" - and minor: "P ==> ~Q" - shows "~P" + and minor: "P \ \ Q" + shows "\ P" by (iprover intro: notI minor major notE) -lemma not_sym: "t ~= s ==> s ~= t" +lemma not_sym: "t \ s \ s \ t" by (erule contrapos_nn) (erule sym) -lemma eq_neq_eq_imp_neq: "[| x = a ; a ~= b; b = y |] ==> x ~= y" +lemma eq_neq_eq_imp_neq: "\x = a; a \ b; b = y\ \ x \ y" by (erule subst, erule ssubst, assumption) subsubsection \Existential quantifier\ -lemma exI: "P x ==> EX x::'a. P x" +lemma exI: "P x \ \x::'a. P x" apply (unfold Ex_def) apply (iprover intro: allI allE impI mp) done lemma exE: - assumes major: "EX x::'a. P(x)" - and minor: "!!x. P(x) ==> Q" + assumes major: "\x::'a. P x" + and minor: "\x. P x \ Q" shows "Q" apply (rule major [unfolded Ex_def, THEN spec, THEN mp]) apply (iprover intro: impI [THEN allI] minor) @@ -432,52 +432,52 @@ subsubsection \Conjunction\ -lemma conjI: "[| P; Q |] ==> P&Q" +lemma conjI: "\P; Q\ \ P \ Q" apply (unfold and_def) apply (iprover intro: impI [THEN allI] mp) done -lemma conjunct1: "[| P & Q |] ==> P" +lemma conjunct1: "\P \ Q\ \ P" apply (unfold and_def) apply (iprover intro: impI dest: spec mp) done -lemma conjunct2: "[| P & Q |] ==> Q" +lemma conjunct2: "\P \ Q\ \ Q" apply (unfold and_def) apply (iprover intro: impI dest: spec mp) done lemma conjE: - assumes major: "P&Q" - and minor: "[| P; Q |] ==> R" - shows "R" + assumes major: "P \ Q" + and minor: "\P; Q\ \ R" + shows R apply (rule minor) apply (rule major [THEN conjunct1]) apply (rule major [THEN conjunct2]) done lemma context_conjI: - assumes "P" "P ==> Q" shows "P & Q" + assumes P "P \ Q" shows "P \ Q" by (iprover intro: conjI assms) subsubsection \Disjunction\ -lemma disjI1: "P ==> P|Q" +lemma disjI1: "P \ P \ Q" apply (unfold or_def) apply (iprover intro: allI impI mp) done -lemma disjI2: "Q ==> P|Q" +lemma disjI2: "Q \ P \ Q" apply (unfold or_def) apply (iprover intro: allI impI mp) done lemma disjE: - assumes major: "P|Q" - and minorP: "P ==> R" - and minorQ: "Q ==> R" - shows "R" + assumes major: "P \ Q" + and minorP: "P \ R" + and minorQ: "Q \ R" + shows R by (iprover intro: minorP minorQ impI major [unfolded or_def, THEN spec, THEN mp, THEN mp]) @@ -485,8 +485,8 @@ subsubsection \Classical logic\ lemma classical: - assumes prem: "~P ==> P" - shows "P" + assumes prem: "\ P \ P" + shows P apply (rule True_or_False [THEN disjE, THEN eqTrueE]) apply assumption apply (rule notI [THEN prem, THEN eqTrueI]) @@ -496,54 +496,54 @@ lemmas ccontr = FalseE [THEN classical] -(*notE with premises exchanged; it discharges ~R so that it can be used to +(*notE with premises exchanged; it discharges \ R so that it can be used to make elimination rules*) lemma rev_notE: - assumes premp: "P" - and premnot: "~R ==> ~P" - shows "R" + assumes premp: P + and premnot: "\ R \ \ P" + shows R apply (rule ccontr) apply (erule notE [OF premnot premp]) done (*Double negation law*) -lemma notnotD: "~~P ==> P" +lemma notnotD: "\\ P \ P" apply (rule classical) apply (erule notE) apply assumption done lemma contrapos_pp: - assumes p1: "Q" - and p2: "~P ==> ~Q" - shows "P" + assumes p1: Q + and p2: "\ P \ \ Q" + shows P by (iprover intro: classical p1 p2 notE) subsubsection \Unique existence\ lemma ex1I: - assumes "P a" "!!x. P(x) ==> x=a" - shows "EX! x. P(x)" + assumes "P a" "\x. P x \ x = a" + shows "\!x. P x" by (unfold Ex1_def, iprover intro: assms exI conjI allI impI) text\Sometimes easier to use: the premises have no shared variables. Safe!\ lemma ex_ex1I: - assumes ex_prem: "EX x. P(x)" - and eq: "!!x y. [| P(x); P(y) |] ==> x=y" - shows "EX! x. P(x)" + assumes ex_prem: "\x. P x" + and eq: "\x y. \P x; P y\ \ x = y" + shows "\!x. P x" by (iprover intro: ex_prem [THEN exE] ex1I eq) lemma ex1E: - assumes major: "EX! x. P(x)" - and minor: "!!x. [| P(x); ALL y. P(y) --> y=x |] ==> R" - shows "R" + assumes major: "\!x. P x" + and minor: "\x. \P x; \y. P y \ y = x\ \ R" + shows R apply (rule major [unfolded Ex1_def, THEN exE]) apply (erule conjE) apply (iprover intro: minor) done -lemma ex1_implies_ex: "EX! x. P x ==> EX x. P x" +lemma ex1_implies_ex: "\!x. P x \ \x. P x" apply (erule ex1E) apply (rule exI) apply assumption @@ -553,59 +553,59 @@ subsubsection \Classical intro rules for disjunction and existential quantifiers\ lemma disjCI: - assumes "~Q ==> P" shows "P|Q" + assumes "\ Q \ P" shows "P \ Q" apply (rule classical) apply (iprover intro: assms disjI1 disjI2 notI elim: notE) done -lemma excluded_middle: "~P | P" +lemma excluded_middle: "\ P \ P" by (iprover intro: disjCI) text \ case distinction as a natural deduction rule. - Note that @{term "~P"} is the second case, not the first + Note that @{term "\ P"} is the second case, not the first \ lemma case_split [case_names True False]: - assumes prem1: "P ==> Q" - and prem2: "~P ==> Q" - shows "Q" + assumes prem1: "P \ Q" + and prem2: "\ P \ Q" + shows Q apply (rule excluded_middle [THEN disjE]) apply (erule prem2) apply (erule prem1) done -(*Classical implies (-->) elimination. *) +(*Classical implies (\) elimination. *) lemma impCE: - assumes major: "P-->Q" - and minor: "~P ==> R" "Q ==> R" - shows "R" + assumes major: "P \ Q" + and minor: "\ P \ R" "Q \ R" + shows R apply (rule excluded_middle [of P, THEN disjE]) apply (iprover intro: minor major [THEN mp])+ done -(*This version of --> elimination works on Q before P. It works best for +(*This version of \ elimination works on Q before P. It works best for those cases in which P holds "almost everywhere". Can't install as default: would break old proofs.*) lemma impCE': - assumes major: "P-->Q" - and minor: "Q ==> R" "~P ==> R" - shows "R" + assumes major: "P \ Q" + and minor: "Q \ R" "\ P \ R" + shows R apply (rule excluded_middle [of P, THEN disjE]) apply (iprover intro: minor major [THEN mp])+ done (*Classical <-> elimination. *) lemma iffCE: - assumes major: "P=Q" - and minor: "[| P; Q |] ==> R" "[| ~P; ~Q |] ==> R" - shows "R" + assumes major: "P = Q" + and minor: "\P; Q\ \ R" "\\ P; \ Q\ \ R" + shows R apply (rule major [THEN iffE]) apply (iprover intro: minor elim: impCE notE) done lemma exCI: - assumes "ALL x. ~P(x) ==> P(a)" - shows "EX x. P(x)" + assumes "\x. \ P x \ P a" + shows "\x. P x" apply (rule ccontr) apply (iprover intro: assms exI allI notI notE [of "\x. P x"]) done @@ -614,9 +614,9 @@ subsubsection \Intuitionistic Reasoning\ lemma impE': - assumes 1: "P --> Q" - and 2: "Q ==> R" - and 3: "P --> Q ==> P" + assumes 1: "P \ Q" + and 2: "Q \ R" + and 3: "P \ Q \ P" shows R proof - from 3 and 1 have P . @@ -625,8 +625,8 @@ qed lemma allE': - assumes 1: "ALL x. P x" - and 2: "P x ==> ALL x. P x ==> Q" + assumes 1: "\x. P x" + and 2: "P x \ \x. P x \ Q" shows Q proof - from 1 have "P x" by (rule spec) @@ -634,16 +634,16 @@ qed lemma notE': - assumes 1: "~ P" - and 2: "~ P ==> P" + assumes 1: "\ P" + and 2: "\ P \ P" shows R proof - from 2 and 1 have P . with 1 show R by (rule notE) qed -lemma TrueE: "True ==> P ==> P" . -lemma notFalseE: "~ False ==> P ==> P" . +lemma TrueE: "True \ P \ P" . +lemma notFalseE: "\ False \ P \ P" . lemmas [Pure.elim!] = disjE iffE FalseE conjE exE TrueE notFalseE and [Pure.intro!] = iffI conjI impI TrueI notI allI refl @@ -660,52 +660,52 @@ axiomatization where eq_reflection: "x = y \ x \ y" (*admissible axiom*) -lemma atomize_all [atomize]: "(!!x. P x) == Trueprop (ALL x. P x)" +lemma atomize_all [atomize]: "(\x. P x) \ Trueprop (\x. P x)" proof - assume "!!x. P x" - then show "ALL x. P x" .. + assume "\x. P x" + then show "\x. P x" .. next - assume "ALL x. P x" - then show "!!x. P x" by (rule allE) + assume "\x. P x" + then show "\x. P x" by (rule allE) qed -lemma atomize_imp [atomize]: "(A ==> B) == Trueprop (A --> B)" +lemma atomize_imp [atomize]: "(A \ B) \ Trueprop (A \ B)" proof - assume r: "A ==> B" - show "A --> B" by (rule impI) (rule r) + assume r: "A \ B" + show "A \ B" by (rule impI) (rule r) next - assume "A --> B" and A + assume "A \ B" and A then show B by (rule mp) qed -lemma atomize_not: "(A ==> False) == Trueprop (~A)" +lemma atomize_not: "(A \ False) \ Trueprop (\ A)" proof - assume r: "A ==> False" - show "~A" by (rule notI) (rule r) + assume r: "A \ False" + show "\ A" by (rule notI) (rule r) next - assume "~A" and A + assume "\ A" and A then show False by (rule notE) qed -lemma atomize_eq [atomize, code]: "(x == y) == Trueprop (x = y)" +lemma atomize_eq [atomize, code]: "(x \ y) \ Trueprop (x = y)" proof - assume "x == y" - show "x = y" by (unfold \x == y\) (rule refl) + assume "x \ y" + show "x = y" by (unfold \x \ y\) (rule refl) next assume "x = y" - then show "x == y" by (rule eq_reflection) + then show "x \ y" by (rule eq_reflection) qed -lemma atomize_conj [atomize]: "(A &&& B) == Trueprop (A & B)" +lemma atomize_conj [atomize]: "(A &&& B) \ Trueprop (A \ B)" proof assume conj: "A &&& B" - show "A & B" + show "A \ B" proof (rule conjI) from conj show A by (rule conjunctionD1) from conj show B by (rule conjunctionD2) qed next - assume conj: "A & B" + assume conj: "A \ B" show "A &&& B" proof - from conj show A .. @@ -719,16 +719,16 @@ subsubsection \Atomizing elimination rules\ -lemma atomize_exL[atomize_elim]: "(!!x. P x ==> Q) == ((EX x. P x) ==> Q)" +lemma atomize_exL[atomize_elim]: "(\x. P x \ Q) \ ((\x. P x) \ Q)" by rule iprover+ -lemma atomize_conjL[atomize_elim]: "(A ==> B ==> C) == (A & B ==> C)" +lemma atomize_conjL[atomize_elim]: "(A \ B \ C) \ (A \ B \ C)" by rule iprover+ -lemma atomize_disjL[atomize_elim]: "((A ==> C) ==> (B ==> C) ==> C) == ((A | B ==> C) ==> C)" +lemma atomize_disjL[atomize_elim]: "((A \ C) \ (B \ C) \ C) \ ((A \ B \ C) \ C)" by rule iprover+ -lemma atomize_elimL[atomize_elim]: "(!!B. (A ==> B) ==> B) == Trueprop A" .. +lemma atomize_elimL[atomize_elim]: "(\B. (A \ B) \ B) \ Trueprop A" .. subsection \Package setup\ @@ -749,14 +749,13 @@ subsubsection \Classical Reasoner setup\ -lemma imp_elim: "P --> Q ==> (~ R ==> P) ==> (Q ==> R) ==> R" +lemma imp_elim: "P \ Q \ (\ R \ P) \ (Q \ R) \ R" by (rule classical) iprover -lemma swap: "~ P ==> (~ R ==> P) ==> R" +lemma swap: "\ P \ (\ R \ P) \ R" by (rule classical) iprover -lemma thin_refl: - "\X. \ x=x; PROP W \ \ PROP W" . +lemma thin_refl: "\X. \x = x; PROP W\ \ PROP W" . ML \ structure Hypsubst = Hypsubst @@ -826,7 +825,7 @@ ML \val HOL_cs = claset_of @{context}\ -lemma contrapos_np: "~ Q ==> (~ P ==> Q) ==> P" +lemma contrapos_np: "\ Q \ (\ P \ Q) \ P" apply (erule swap) apply (erule (1) meta_mp) done @@ -871,83 +870,83 @@ lemma the_equality [intro]: assumes "P a" - and "!!x. P x ==> x=a" + and "\x. P x \ x = a" shows "(THE x. P x) = a" by (blast intro: assms trans [OF arg_cong [where f=The] the_eq_trivial]) lemma theI: - assumes "P a" and "!!x. P x ==> x=a" + assumes "P a" and "\x. P x \ x = a" shows "P (THE x. P x)" by (iprover intro: assms the_equality [THEN ssubst]) -lemma theI': "EX! x. P x ==> P (THE x. P x)" +lemma theI': "\!x. P x \ P (THE x. P x)" by (blast intro: theI) (*Easier to apply than theI: only one occurrence of P*) lemma theI2: - assumes "P a" "!!x. P x ==> x=a" "!!x. P x ==> Q x" + assumes "P a" "\x. P x \ x = a" "\x. P x \ Q x" shows "Q (THE x. P x)" by (iprover intro: assms theI) -lemma the1I2: assumes "EX! x. P x" "\x. P x \ Q x" shows "Q (THE x. P x)" +lemma the1I2: assumes "\!x. P x" "\x. P x \ Q x" shows "Q (THE x. P x)" by(iprover intro:assms(2) theI2[where P=P and Q=Q] ex1E[OF assms(1)] elim:allE impE) -lemma the1_equality [elim?]: "[| EX!x. P x; P a |] ==> (THE x. P x) = a" +lemma the1_equality [elim?]: "\\!x. P x; P a\ \ (THE x. P x) = a" by blast -lemma the_sym_eq_trivial: "(THE y. x=y) = x" +lemma the_sym_eq_trivial: "(THE y. x = y) = x" by blast subsubsection \Simplifier\ -lemma eta_contract_eq: "(%s. f s) = f" .. +lemma eta_contract_eq: "(\s. f s) = f" .. lemma simp_thms: - shows not_not: "(~ ~ P) = P" - and Not_eq_iff: "((~P) = (~Q)) = (P = Q)" + shows not_not: "(\ \ P) = P" + and Not_eq_iff: "((\ P) = (\ Q)) = (P = Q)" and - "(P ~= Q) = (P = (~Q))" - "(P | ~P) = True" "(~P | P) = True" + "(P \ Q) = (P = (\ Q))" + "(P \ \P) = True" "(\ P \ P) = True" "(x = x) = True" and not_True_eq_False [code]: "(\ True) = False" and not_False_eq_True [code]: "(\ False) = True" and - "(~P) ~= P" "P ~= (~P)" - "(True=P) = P" + "(\ P) \ P" "P \ (\ P)" + "(True = P) = P" and eq_True: "(P = True) = P" - and "(False=P) = (~P)" + and "(False = P) = (\ P)" and eq_False: "(P = False) = (\ P)" and - "(True --> P) = P" "(False --> P) = True" - "(P --> True) = True" "(P --> P) = True" - "(P --> False) = (~P)" "(P --> ~P) = (~P)" - "(P & True) = P" "(True & P) = P" - "(P & False) = False" "(False & P) = False" - "(P & P) = P" "(P & (P & Q)) = (P & Q)" - "(P & ~P) = False" "(~P & P) = False" - "(P | True) = True" "(True | P) = True" - "(P | False) = P" "(False | P) = P" - "(P | P) = P" "(P | (P | Q)) = (P | Q)" and - "(ALL x. P) = P" "(EX x. P) = P" "EX x. x=t" "EX x. t=x" + "(True \ P) = P" "(False \ P) = True" + "(P \ True) = True" "(P \ P) = True" + "(P \ False) = (\ P)" "(P \ \ P) = (\ P)" + "(P \ True) = P" "(True \ P) = P" + "(P \ False) = False" "(False \ P) = False" + "(P \ P) = P" "(P \ (P \ Q)) = (P \ Q)" + "(P \ \ P) = False" "(\ P \ P) = False" + "(P \ True) = True" "(True \ P) = True" + "(P \ False) = P" "(False \ P) = P" + "(P \ P) = P" "(P \ (P \ Q)) = (P \ Q)" and + "(\x. P) = P" "(\x. P) = P" "\x. x = t" "\x. t = x" and - "!!P. (EX x. x=t & P(x)) = P(t)" - "!!P. (EX x. t=x & P(x)) = P(t)" - "!!P. (ALL x. x=t --> P(x)) = P(t)" - "!!P. (ALL x. t=x --> P(x)) = P(t)" + "\P. (\x. x = t \ P x) = P t" + "\P. (\x. t = x \ P x) = P t" + "\P. (\x. x = t \ P x) = P t" + "\P. (\x. t = x \ P x) = P t" by (blast, blast, blast, blast, blast, iprover+) -lemma disj_absorb: "(A | A) = A" +lemma disj_absorb: "(A \ A) = A" by blast -lemma disj_left_absorb: "(A | (A | B)) = (A | B)" +lemma disj_left_absorb: "(A \ (A \ B)) = (A \ B)" by blast -lemma conj_absorb: "(A & A) = A" +lemma conj_absorb: "(A \ A) = A" by blast -lemma conj_left_absorb: "(A & (A & B)) = (A & B)" +lemma conj_left_absorb: "(A \ (A \ B)) = (A \ B)" by blast lemma eq_ac: @@ -957,83 +956,83 @@ lemma neq_commute: "a \ b \ b \ a" by iprover lemma conj_comms: - shows conj_commute: "(P&Q) = (Q&P)" - and conj_left_commute: "(P&(Q&R)) = (Q&(P&R))" by iprover+ -lemma conj_assoc: "((P&Q)&R) = (P&(Q&R))" by iprover + shows conj_commute: "(P \ Q) = (Q \ P)" + and conj_left_commute: "(P \ (Q \ R)) = (Q \ (P \ R))" by iprover+ +lemma conj_assoc: "((P \ Q) \ R) = (P \ (Q \ R))" by iprover lemmas conj_ac = conj_commute conj_left_commute conj_assoc lemma disj_comms: - shows disj_commute: "(P|Q) = (Q|P)" - and disj_left_commute: "(P|(Q|R)) = (Q|(P|R))" by iprover+ -lemma disj_assoc: "((P|Q)|R) = (P|(Q|R))" by iprover + shows disj_commute: "(P \ Q) = (Q \ P)" + and disj_left_commute: "(P \ (Q \ R)) = (Q \ (P \ R))" by iprover+ +lemma disj_assoc: "((P \ Q) \ R) = (P \ (Q \ R))" by iprover lemmas disj_ac = disj_commute disj_left_commute disj_assoc -lemma conj_disj_distribL: "(P&(Q|R)) = (P&Q | P&R)" by iprover -lemma conj_disj_distribR: "((P|Q)&R) = (P&R | Q&R)" by iprover +lemma conj_disj_distribL: "(P \ (Q \ R)) = (P \ Q \ P \ R)" by iprover +lemma conj_disj_distribR: "((P \ Q) \ R) = (P \ R \ Q \ R)" by iprover -lemma disj_conj_distribL: "(P|(Q&R)) = ((P|Q) & (P|R))" by iprover -lemma disj_conj_distribR: "((P&Q)|R) = ((P|R) & (Q|R))" by iprover +lemma disj_conj_distribL: "(P \ (Q \ R)) = ((P \ Q) \ (P \ R))" by iprover +lemma disj_conj_distribR: "((P \ Q) \ R) = ((P \ R) \ (Q \ R))" by iprover -lemma imp_conjR: "(P --> (Q&R)) = ((P-->Q) & (P-->R))" by iprover -lemma imp_conjL: "((P&Q) -->R) = (P --> (Q --> R))" by iprover -lemma imp_disjL: "((P|Q) --> R) = ((P-->R)&(Q-->R))" by iprover +lemma imp_conjR: "(P \ (Q \ R)) = ((P \ Q) \ (P \ R))" by iprover +lemma imp_conjL: "((P \ Q) \ R) = (P \ (Q \ R))" by iprover +lemma imp_disjL: "((P \ Q) \ R) = ((P \ R) \ (Q \ R))" by iprover text \These two are specialized, but @{text imp_disj_not1} is useful in @{text "Auth/Yahalom"}.\ -lemma imp_disj_not1: "(P --> Q | R) = (~Q --> P --> R)" by blast -lemma imp_disj_not2: "(P --> Q | R) = (~R --> P --> Q)" by blast +lemma imp_disj_not1: "(P \ Q \ R) = (\ Q \ P \ R)" by blast +lemma imp_disj_not2: "(P \ Q \ R) = (\ R \ P \ Q)" by blast -lemma imp_disj1: "((P-->Q)|R) = (P--> Q|R)" by blast -lemma imp_disj2: "(Q|(P-->R)) = (P--> Q|R)" by blast +lemma imp_disj1: "((P \ Q) \ R) = (P \ Q \ R)" by blast +lemma imp_disj2: "(Q \ (P \ R)) = (P \ Q \ R)" by blast -lemma imp_cong: "(P = P') ==> (P' ==> (Q = Q')) ==> ((P --> Q) = (P' --> Q'))" +lemma imp_cong: "(P = P') \ (P' \ (Q = Q')) \ ((P \ Q) = (P' \ Q'))" by iprover -lemma de_Morgan_disj: "(~(P | Q)) = (~P & ~Q)" by iprover -lemma de_Morgan_conj: "(~(P & Q)) = (~P | ~Q)" by blast -lemma not_imp: "(~(P --> Q)) = (P & ~Q)" by blast -lemma not_iff: "(P~=Q) = (P = (~Q))" by blast -lemma disj_not1: "(~P | Q) = (P --> Q)" by blast -lemma disj_not2: "(P | ~Q) = (Q --> P)" -- \changes orientation :-(\ +lemma de_Morgan_disj: "(\ (P \ Q)) = (\ P \ \ Q)" by iprover +lemma de_Morgan_conj: "(\ (P \ Q)) = (\ P \ \ Q)" by blast +lemma not_imp: "(\ (P \ Q)) = (P \ \ Q)" by blast +lemma not_iff: "(P \ Q) = (P = (\ Q))" by blast +lemma disj_not1: "(\ P \ Q) = (P \ Q)" by blast +lemma disj_not2: "(P \ \ Q) = (Q \ P)" -- \changes orientation :-(\ by blast -lemma imp_conv_disj: "(P --> Q) = ((~P) | Q)" by blast +lemma imp_conv_disj: "(P \ Q) = ((\ P) \ Q)" by blast -lemma iff_conv_conj_imp: "(P = Q) = ((P --> Q) & (Q --> P))" by iprover +lemma iff_conv_conj_imp: "(P = Q) = ((P \ Q) \ (Q \ P))" by iprover -lemma cases_simp: "((P --> Q) & (~P --> Q)) = Q" +lemma cases_simp: "((P \ Q) \ (\ P \ Q)) = Q" -- \Avoids duplication of subgoals after @{text split_if}, when the true and false\ -- \cases boil down to the same thing.\ by blast -lemma not_all: "(~ (! x. P(x))) = (? x.~P(x))" by blast -lemma imp_all: "((! x. P x) --> Q) = (? x. P x --> Q)" by blast -lemma not_ex: "(~ (? x. P(x))) = (! x.~P(x))" by iprover -lemma imp_ex: "((? x. P x) --> Q) = (! x. P x --> Q)" by iprover -lemma all_not_ex: "(ALL x. P x) = (~ (EX x. ~ P x ))" by blast +lemma not_all: "(\ (\x. P x)) = (\x. \ P x)" by blast +lemma imp_all: "((\x. P x) \ Q) = (\x. P x \ Q)" by blast +lemma not_ex: "(\ (\x. P x)) = (\x. \ P x)" by iprover +lemma imp_ex: "((\x. P x) \ Q) = (\x. P x \ Q)" by iprover +lemma all_not_ex: "(\x. P x) = (\ (\x. \ P x ))" by blast declare All_def [no_atp] -lemma ex_disj_distrib: "(? x. P(x) | Q(x)) = ((? x. P(x)) | (? x. Q(x)))" by iprover -lemma all_conj_distrib: "(!x. P(x) & Q(x)) = ((! x. P(x)) & (! x. Q(x)))" by iprover +lemma ex_disj_distrib: "(\x. P x \ Q x) = ((\x. P x) \ (\x. Q x))" by iprover +lemma all_conj_distrib: "(\x. P x \ Q x) = ((\x. P x) \ (\x. Q x))" by iprover text \ - \medskip The @{text "&"} congruence rule: not included by default! + \medskip The @{text "\"} congruence rule: not included by default! May slow rewrite proofs down by as much as 50\%\ lemma conj_cong: - "(P = P') ==> (P' ==> (Q = Q')) ==> ((P & Q) = (P' & Q'))" + "(P = P') \ (P' \ (Q = Q')) \ ((P \ Q) = (P' \ Q'))" by iprover lemma rev_conj_cong: - "(Q = Q') ==> (Q' ==> (P = P')) ==> ((P & Q) = (P' & Q'))" + "(Q = Q') \ (Q' \ (P = P')) \ ((P \ Q) = (P' \ Q'))" by iprover text \The @{text "|"} congruence rule: not included by default!\ lemma disj_cong: - "(P = P') ==> (~P' ==> (Q = Q')) ==> ((P | Q) = (P' | Q'))" + "(P = P') \ (\ P' \ (Q = Q')) \ ((P \ Q) = (P' \ Q'))" by blast @@ -1045,19 +1044,19 @@ lemma if_False [code]: "(if False then x else y) = y" by (unfold If_def) blast -lemma if_P: "P ==> (if P then x else y) = x" +lemma if_P: "P \ (if P then x else y) = x" by (unfold If_def) blast -lemma if_not_P: "~P ==> (if P then x else y) = y" +lemma if_not_P: "\ P \ (if P then x else y) = y" by (unfold If_def) blast -lemma split_if: "P (if Q then x else y) = ((Q --> P(x)) & (~Q --> P(y)))" +lemma split_if: "P (if Q then x else y) = ((Q \ P x) \ (\ Q \ P y))" apply (rule case_split [of Q]) apply (simplesubst if_P) prefer 3 apply (simplesubst if_not_P, blast+) done -lemma split_if_asm: "P (if Q then x else y) = (~((Q & ~P x) | (~Q & ~P y)))" +lemma split_if_asm: "P (if Q then x else y) = (\ ((Q \ \ P x) \ (\ Q \ \ P y)))" by (simplesubst split_if, blast) lemmas if_splits [no_atp] = split_if split_if_asm @@ -1068,24 +1067,23 @@ lemma if_eq_cancel: "(if x = y then y else x) = x" by (simplesubst split_if, blast) -lemma if_bool_eq_conj: -"(if P then Q else R) = ((P-->Q) & (~P-->R))" - -- \This form is useful for expanding @{text "if"}s on the RIGHT of the @{text "==>"} symbol.\ +lemma if_bool_eq_conj: "(if P then Q else R) = ((P \ Q) \ (\ P \ R))" + -- \This form is useful for expanding @{text "if"}s on the RIGHT of the @{text "\"} symbol.\ by (rule split_if) -lemma if_bool_eq_disj: "(if P then Q else R) = ((P&Q) | (~P&R))" +lemma if_bool_eq_disj: "(if P then Q else R) = ((P \ Q) \ (\ P \ R))" -- \And this form is useful for expanding @{text "if"}s on the LEFT.\ by (simplesubst split_if) blast -lemma Eq_TrueI: "P ==> P == True" by (unfold atomize_eq) iprover -lemma Eq_FalseI: "~P ==> P == False" by (unfold atomize_eq) iprover +lemma Eq_TrueI: "P \ P \ True" by (unfold atomize_eq) iprover +lemma Eq_FalseI: "\ P \ P \ False" by (unfold atomize_eq) iprover text \\medskip let rules for simproc\ -lemma Let_folded: "f x \ g x \ Let x f \ Let x g" +lemma Let_folded: "f x \ g x \ Let x f \ Let x g" by (unfold Let_def) -lemma Let_unfold: "f x \ g \ Let x f \ g" +lemma Let_unfold: "f x \ g \ Let x f \ g" by (unfold Let_def) text \ @@ -1094,8 +1092,8 @@ its premise. \ -definition simp_implies :: "[prop, prop] => prop" (infixr "=simp=>" 1) where - "simp_implies \ op ==>" +definition simp_implies :: "[prop, prop] \ prop" (infixr "=simp=>" 1) where + "simp_implies \ op \" lemma simp_impliesI: assumes PQ: "(PROP P \ PROP Q)" @@ -1116,9 +1114,9 @@ done lemma simp_implies_cong: - assumes PP' :"PROP P == PROP P'" - and P'QQ': "PROP P' ==> (PROP Q == PROP Q')" - shows "(PROP P =simp=> PROP Q) == (PROP P' =simp=> PROP Q')" + assumes PP' :"PROP P \ PROP P'" + and P'QQ': "PROP P' \ (PROP Q \ PROP Q')" + shows "(PROP P =simp=> PROP Q) \ (PROP P' =simp=> PROP Q')" proof (unfold simp_implies_def, rule equal_intr_rule) assume PQ: "PROP P \ PROP Q" and P': "PROP P'" @@ -1166,10 +1164,10 @@ Simplifier.method_setup Splitter.split_modifiers \ -simproc_setup defined_Ex ("EX x. P x") = \fn _ => Quantifier1.rearrange_ex\ -simproc_setup defined_All ("ALL x. P x") = \fn _ => Quantifier1.rearrange_all\ +simproc_setup defined_Ex ("\x. P x") = \fn _ => Quantifier1.rearrange_ex\ +simproc_setup defined_All ("\x. P x") = \fn _ => Quantifier1.rearrange_all\ -text \Simproc for proving @{text "(y = x) == False"} from premise @{text "~(x = y)"}:\ +text \Simproc for proving @{text "(y = x) \ False"} from premise @{text "\ (x = y)"}:\ simproc_setup neq ("x = y") = \fn _ => let @@ -1277,22 +1275,22 @@ by(rule swap_prems_eq) lemma ex_simps: - "!!P Q. (EX x. P x & Q) = ((EX x. P x) & Q)" - "!!P Q. (EX x. P & Q x) = (P & (EX x. Q x))" - "!!P Q. (EX x. P x | Q) = ((EX x. P x) | Q)" - "!!P Q. (EX x. P | Q x) = (P | (EX x. Q x))" - "!!P Q. (EX x. P x --> Q) = ((ALL x. P x) --> Q)" - "!!P Q. (EX x. P --> Q x) = (P --> (EX x. Q x))" + "\P Q. (\x. P x \ Q) = ((\x. P x) \ Q)" + "\P Q. (\x. P \ Q x) = (P \ (\x. Q x))" + "\P Q. (\x. P x \ Q) = ((\x. P x) \ Q)" + "\P Q. (\x. P \ Q x) = (P \ (\x. Q x))" + "\P Q. (\x. P x \ Q) = ((\x. P x) \ Q)" + "\P Q. (\x. P \ Q x) = (P \ (\x. Q x))" -- \Miniscoping: pushing in existential quantifiers.\ by (iprover | blast)+ lemma all_simps: - "!!P Q. (ALL x. P x & Q) = ((ALL x. P x) & Q)" - "!!P Q. (ALL x. P & Q x) = (P & (ALL x. Q x))" - "!!P Q. (ALL x. P x | Q) = ((ALL x. P x) | Q)" - "!!P Q. (ALL x. P | Q x) = (P | (ALL x. Q x))" - "!!P Q. (ALL x. P x --> Q) = ((EX x. P x) --> Q)" - "!!P Q. (ALL x. P --> Q x) = (P --> (ALL x. Q x))" + "\P Q. (\x. P x \ Q) = ((\x. P x) \ Q)" + "\P Q. (\x. P \ Q x) = (P \ (\x. Q x))" + "\P Q. (\x. P x \ Q) = ((\x. P x) \ Q)" + "\P Q. (\x. P \ Q x) = (P \ (\x. Q x))" + "\P Q. (\x. P x \ Q) = ((\x. P x) \ Q)" + "\P Q. (\x. P \ Q x) = (P \ (\x. Q x))" -- \Miniscoping: pushing in universal quantifiers.\ by (iprover | blast)+ @@ -1308,7 +1306,7 @@ (*In general it seems wrong to add distributive laws by default: they might cause exponential blow-up. But imp_disjL has been in for a while and cannot be removed without affecting existing proofs. Moreover, - rewriting by "(P|Q --> R) = ((P-->R)&(Q-->R))" might be justified on the + rewriting by "(P \ Q \ R) = ((P \ R) \ (Q \ R))" might be justified on the grounds that it allows simplification of R in the two cases.*) conj_assoc disj_assoc @@ -1332,7 +1330,7 @@ ML \val HOL_ss = simpset_of @{context}\ -text \Simplifies x assuming c and y assuming ~c\ +text \Simplifies x assuming c and y assuming \ c\ lemma if_cong: assumes "b = c" and "c \ x = u" @@ -1458,13 +1456,13 @@ fn _ => Induct.map_simpset (fn ss => ss addsimprocs [Simplifier.simproc_global @{theory} "swap_induct_false" - ["induct_false ==> PROP P ==> PROP Q"] + ["induct_false \ PROP P \ PROP Q"] (fn _ => (fn _ $ (P as _ $ @{const induct_false}) $ (_ $ Q $ _) => if P <> Q then SOME Drule.swap_prems_eq else NONE | _ => NONE)), Simplifier.simproc_global @{theory} "induct_equal_conj_curry" - ["induct_conj P Q ==> PROP R"] + ["induct_conj P Q \ PROP R"] (fn _ => (fn _ $ (_ $ P) $ _ => let @@ -1589,19 +1587,19 @@ subsection \Other simple lemmas and lemma duplicates\ -lemma ex1_eq [iff]: "EX! x. x = t" "EX! x. t = x" +lemma ex1_eq [iff]: "\!x. x = t" "\!x. t = x" by blast+ -lemma choice_eq: "(ALL x. EX! y. P x y) = (EX! f. ALL x. P x (f x))" +lemma choice_eq: "(\x. \!y. P x y) = (\!f. \x. P x (f x))" apply (rule iffI) - apply (rule_tac a = "%x. THE y. P x y" in ex1I) + apply (rule_tac a = "\x. THE y. P x y" in ex1I) apply (fast dest!: theI') apply (fast intro: the1_equality [symmetric]) apply (erule ex1E) apply (rule allI) apply (rule ex1I) apply (erule spec) - apply (erule_tac x = "%z. if z = x then y else f z" in allE) + apply (erule_tac x = "\z. if z = x then y else f z" in allE) apply (erule impE) apply (rule allI) apply (case_tac "xa = x") diff -r d8d85a8172b5 -r 36d9f215c982 src/HOL/Tools/cnf.ML --- a/src/HOL/Tools/cnf.ML Sat Jul 18 22:58:50 2015 +0200 +++ b/src/HOL/Tools/cnf.ML Sun Jul 19 00:03:10 2015 +0200 @@ -54,41 +54,41 @@ structure CNF : CNF = struct -val clause2raw_notE = @{lemma "[| P; ~P |] ==> False" by auto}; -val clause2raw_not_disj = @{lemma "[| ~P; ~Q |] ==> ~(P | Q)" by auto}; -val clause2raw_not_not = @{lemma "P ==> ~~P" by auto}; +val clause2raw_notE = @{lemma "\P; \P\ \ False" by auto}; +val clause2raw_not_disj = @{lemma "\\ P; \ Q\ \ \ (P \ Q)" by auto}; +val clause2raw_not_not = @{lemma "P \ \\ P" by auto}; val iff_refl = @{lemma "(P::bool) = P" by auto}; val iff_trans = @{lemma "[| (P::bool) = Q; Q = R |] ==> P = R" by auto}; -val conj_cong = @{lemma "[| P = P'; Q = Q' |] ==> (P & Q) = (P' & Q')" by auto}; -val disj_cong = @{lemma "[| P = P'; Q = Q' |] ==> (P | Q) = (P' | Q')" by auto}; +val conj_cong = @{lemma "[| P = P'; Q = Q' |] ==> (P \ Q) = (P' \ Q')" by auto}; +val disj_cong = @{lemma "[| P = P'; Q = Q' |] ==> (P \ Q) = (P' \ Q')" by auto}; -val make_nnf_imp = @{lemma "[| (~P) = P'; Q = Q' |] ==> (P --> Q) = (P' | Q')" by auto}; -val make_nnf_iff = @{lemma "[| P = P'; (~P) = NP; Q = Q'; (~Q) = NQ |] ==> (P = Q) = ((P' | NQ) & (NP | Q'))" by auto}; +val make_nnf_imp = @{lemma "[| (~P) = P'; Q = Q' |] ==> (P \ Q) = (P' \ Q')" by auto}; +val make_nnf_iff = @{lemma "[| P = P'; (~P) = NP; Q = Q'; (~Q) = NQ |] ==> (P = Q) = ((P' \ NQ) \ (NP \ Q'))" by auto}; val make_nnf_not_false = @{lemma "(~False) = True" by auto}; val make_nnf_not_true = @{lemma "(~True) = False" by auto}; -val make_nnf_not_conj = @{lemma "[| (~P) = P'; (~Q) = Q' |] ==> (~(P & Q)) = (P' | Q')" by auto}; -val make_nnf_not_disj = @{lemma "[| (~P) = P'; (~Q) = Q' |] ==> (~(P | Q)) = (P' & Q')" by auto}; -val make_nnf_not_imp = @{lemma "[| P = P'; (~Q) = Q' |] ==> (~(P --> Q)) = (P' & Q')" by auto}; -val make_nnf_not_iff = @{lemma "[| P = P'; (~P) = NP; Q = Q'; (~Q) = NQ |] ==> (~(P = Q)) = ((P' | Q') & (NP | NQ))" by auto}; +val make_nnf_not_conj = @{lemma "[| (~P) = P'; (~Q) = Q' |] ==> (~(P \ Q)) = (P' \ Q')" by auto}; +val make_nnf_not_disj = @{lemma "[| (~P) = P'; (~Q) = Q' |] ==> (~(P \ Q)) = (P' \ Q')" by auto}; +val make_nnf_not_imp = @{lemma "[| P = P'; (~Q) = Q' |] ==> (~(P \ Q)) = (P' \ Q')" by auto}; +val make_nnf_not_iff = @{lemma "[| P = P'; (~P) = NP; Q = Q'; (~Q) = NQ |] ==> (~(P = Q)) = ((P' \ Q') \ (NP \ NQ))" by auto}; val make_nnf_not_not = @{lemma "P = P' ==> (~~P) = P'" by auto}; -val simp_TF_conj_True_l = @{lemma "[| P = True; Q = Q' |] ==> (P & Q) = Q'" by auto}; -val simp_TF_conj_True_r = @{lemma "[| P = P'; Q = True |] ==> (P & Q) = P'" by auto}; -val simp_TF_conj_False_l = @{lemma "P = False ==> (P & Q) = False" by auto}; -val simp_TF_conj_False_r = @{lemma "Q = False ==> (P & Q) = False" by auto}; -val simp_TF_disj_True_l = @{lemma "P = True ==> (P | Q) = True" by auto}; -val simp_TF_disj_True_r = @{lemma "Q = True ==> (P | Q) = True" by auto}; -val simp_TF_disj_False_l = @{lemma "[| P = False; Q = Q' |] ==> (P | Q) = Q'" by auto}; -val simp_TF_disj_False_r = @{lemma "[| P = P'; Q = False |] ==> (P | Q) = P'" by auto}; +val simp_TF_conj_True_l = @{lemma "[| P = True; Q = Q' |] ==> (P \ Q) = Q'" by auto}; +val simp_TF_conj_True_r = @{lemma "[| P = P'; Q = True |] ==> (P \ Q) = P'" by auto}; +val simp_TF_conj_False_l = @{lemma "P = False ==> (P \ Q) = False" by auto}; +val simp_TF_conj_False_r = @{lemma "Q = False ==> (P \ Q) = False" by auto}; +val simp_TF_disj_True_l = @{lemma "P = True ==> (P \ Q) = True" by auto}; +val simp_TF_disj_True_r = @{lemma "Q = True ==> (P \ Q) = True" by auto}; +val simp_TF_disj_False_l = @{lemma "[| P = False; Q = Q' |] ==> (P \ Q) = Q'" by auto}; +val simp_TF_disj_False_r = @{lemma "[| P = P'; Q = False |] ==> (P \ Q) = P'" by auto}; -val make_cnf_disj_conj_l = @{lemma "[| (P | R) = PR; (Q | R) = QR |] ==> ((P & Q) | R) = (PR & QR)" by auto}; -val make_cnf_disj_conj_r = @{lemma "[| (P | Q) = PQ; (P | R) = PR |] ==> (P | (Q & R)) = (PQ & PR)" by auto}; +val make_cnf_disj_conj_l = @{lemma "[| (P \ R) = PR; (Q \ R) = QR |] ==> ((P \ Q) \ R) = (PR \ QR)" by auto}; +val make_cnf_disj_conj_r = @{lemma "[| (P \ Q) = PQ; (P \ R) = PR |] ==> (P \ (Q \ R)) = (PQ \ PR)" by auto}; -val make_cnfx_disj_ex_l = @{lemma "((EX (x::bool). P x) | Q) = (EX x. P x | Q)" by auto}; -val make_cnfx_disj_ex_r = @{lemma "(P | (EX (x::bool). Q x)) = (EX x. P | Q x)" by auto}; -val make_cnfx_newlit = @{lemma "(P | Q) = (EX x. (P | x) & (Q | ~x))" by auto}; -val make_cnfx_ex_cong = @{lemma "(ALL (x::bool). P x = Q x) ==> (EX x. P x) = (EX x. Q x)" by auto}; +val make_cnfx_disj_ex_l = @{lemma "((EX (x::bool). P x) \ Q) = (EX x. P x \ Q)" by auto}; +val make_cnfx_disj_ex_r = @{lemma "(P \ (EX (x::bool). Q x)) = (EX x. P \ Q x)" by auto}; +val make_cnfx_newlit = @{lemma "(P \ Q) = (EX x. (P \ x) \ (Q \ ~x))" by auto}; +val make_cnfx_ex_cong = @{lemma "(ALL (x::bool). P x = Q x) \ (EX x. P x) = (EX x. Q x)" by auto}; val weakening_thm = @{lemma "[| P; Q |] ==> Q" by auto};