--- a/src/FOL/FOL.thy Thu Jan 03 21:48:05 2019 +0100
+++ b/src/FOL/FOL.thy Thu Jan 03 22:19:19 2019 +0100
@@ -17,18 +17,18 @@
subsection \<open>The classical axiom\<close>
axiomatization where
- classical: "(\<not> P \<Longrightarrow> P) \<Longrightarrow> P"
+ classical: \<open>(\<not> P \<Longrightarrow> P) \<Longrightarrow> P\<close>
subsection \<open>Lemmas and proof tools\<close>
-lemma ccontr: "(\<not> P \<Longrightarrow> False) \<Longrightarrow> P"
+lemma ccontr: \<open>(\<not> P \<Longrightarrow> False) \<Longrightarrow> P\<close>
by (erule FalseE [THEN classical])
subsubsection \<open>Classical introduction rules for \<open>\<or>\<close> and \<open>\<exists>\<close>\<close>
-lemma disjCI: "(\<not> Q \<Longrightarrow> P) \<Longrightarrow> P \<or> Q"
+lemma disjCI: \<open>(\<not> Q \<Longrightarrow> P) \<Longrightarrow> P \<or> Q\<close>
apply (rule classical)
apply (assumption | erule meta_mp | rule disjI1 notI)+
apply (erule notE disjI2)+
@@ -36,31 +36,31 @@
text \<open>Introduction rule involving only \<open>\<exists>\<close>\<close>
lemma ex_classical:
- assumes r: "\<not> (\<exists>x. P(x)) \<Longrightarrow> P(a)"
- shows "\<exists>x. P(x)"
+ assumes r: \<open>\<not> (\<exists>x. P(x)) \<Longrightarrow> P(a)\<close>
+ shows \<open>\<exists>x. P(x)\<close>
apply (rule classical)
apply (rule exI, erule r)
done
text \<open>Version of above, simplifying \<open>\<not>\<exists>\<close> to \<open>\<forall>\<not>\<close>.\<close>
lemma exCI:
- assumes r: "\<forall>x. \<not> P(x) \<Longrightarrow> P(a)"
- shows "\<exists>x. P(x)"
+ assumes r: \<open>\<forall>x. \<not> P(x) \<Longrightarrow> P(a)\<close>
+ shows \<open>\<exists>x. P(x)\<close>
apply (rule ex_classical)
apply (rule notI [THEN allI, THEN r])
apply (erule notE)
apply (erule exI)
done
-lemma excluded_middle: "\<not> P \<or> P"
+lemma excluded_middle: \<open>\<not> P \<or> P\<close>
apply (rule disjCI)
apply assumption
done
lemma case_split [case_names True False]:
- assumes r1: "P \<Longrightarrow> Q"
- and r2: "\<not> P \<Longrightarrow> Q"
- shows Q
+ assumes r1: \<open>P \<Longrightarrow> Q\<close>
+ and r2: \<open>\<not> P \<Longrightarrow> Q\<close>
+ shows \<open>Q\<close>
apply (rule excluded_middle [THEN disjE])
apply (erule r2)
apply (erule r1)
@@ -81,10 +81,10 @@
text \<open>Classical implies (\<open>\<longrightarrow>\<close>) elimination.\<close>
lemma impCE:
- assumes major: "P \<longrightarrow> Q"
- and r1: "\<not> P \<Longrightarrow> R"
- and r2: "Q \<Longrightarrow> R"
- shows R
+ assumes major: \<open>P \<longrightarrow> Q\<close>
+ and r1: \<open>\<not> P \<Longrightarrow> R\<close>
+ and r2: \<open>Q \<Longrightarrow> R\<close>
+ shows \<open>R\<close>
apply (rule excluded_middle [THEN disjE])
apply (erule r1)
apply (rule r2)
@@ -96,10 +96,10 @@
Can't install as default: would break old proofs.
\<close>
lemma impCE':
- assumes major: "P \<longrightarrow> Q"
- and r1: "Q \<Longrightarrow> R"
- and r2: "\<not> P \<Longrightarrow> R"
- shows R
+ assumes major: \<open>P \<longrightarrow> Q\<close>
+ and r1: \<open>Q \<Longrightarrow> R\<close>
+ and r2: \<open>\<not> P \<Longrightarrow> R\<close>
+ shows \<open>R\<close>
apply (rule excluded_middle [THEN disjE])
apply (erule r2)
apply (rule r1)
@@ -107,13 +107,13 @@
done
text \<open>Double negation law.\<close>
-lemma notnotD: "\<not> \<not> P \<Longrightarrow> P"
+lemma notnotD: \<open>\<not> \<not> P \<Longrightarrow> P\<close>
apply (rule classical)
apply (erule notE)
apply assumption
done
-lemma contrapos2: "\<lbrakk>Q; \<not> P \<Longrightarrow> \<not> Q\<rbrakk> \<Longrightarrow> P"
+lemma contrapos2: \<open>\<lbrakk>Q; \<not> P \<Longrightarrow> \<not> Q\<rbrakk> \<Longrightarrow> P\<close>
apply (rule classical)
apply (drule (1) meta_mp)
apply (erule (1) notE)
@@ -127,10 +127,10 @@
\<open>\<not> P \<Longrightarrow> \<not> Q\<close> and \<open>P \<Longrightarrow> Q\<close>.
\<close>
lemma iffCE:
- assumes major: "P \<longleftrightarrow> Q"
- and r1: "\<lbrakk>P; Q\<rbrakk> \<Longrightarrow> R"
- and r2: "\<lbrakk>\<not> P; \<not> Q\<rbrakk> \<Longrightarrow> R"
- shows R
+ assumes major: \<open>P \<longleftrightarrow> Q\<close>
+ and r1: \<open>\<lbrakk>P; Q\<rbrakk> \<Longrightarrow> R\<close>
+ and r2: \<open>\<lbrakk>\<not> P; \<not> Q\<rbrakk> \<Longrightarrow> R\<close>
+ shows \<open>R\<close>
apply (rule major [unfolded iff_def, THEN conjE])
apply (elim impCE)
apply (erule (1) r2)
@@ -141,32 +141,32 @@
(*Better for fast_tac: needs no quantifier duplication!*)
lemma alt_ex1E:
- assumes major: "\<exists>! x. P(x)"
- and r: "\<And>x. \<lbrakk>P(x); \<forall>y y'. P(y) \<and> P(y') \<longrightarrow> y = y'\<rbrakk> \<Longrightarrow> R"
- shows R
+ assumes major: \<open>\<exists>! x. P(x)\<close>
+ and r: \<open>\<And>x. \<lbrakk>P(x); \<forall>y y'. P(y) \<and> P(y') \<longrightarrow> y = y'\<rbrakk> \<Longrightarrow> R\<close>
+ shows \<open>R\<close>
using major
proof (rule ex1E)
fix x
- assume * : "\<forall>y. P(y) \<longrightarrow> y = x"
- assume "P(x)"
- then show R
+ assume * : \<open>\<forall>y. P(y) \<longrightarrow> y = x\<close>
+ assume \<open>P(x)\<close>
+ then show \<open>R\<close>
proof (rule r)
{
fix y y'
- assume "P(y)" and "P(y')"
- with * have "x = y" and "x = y'"
+ assume \<open>P(y)\<close> and \<open>P(y')\<close>
+ with * have \<open>x = y\<close> and \<open>x = y'\<close>
by - (tactic "IntPr.fast_tac @{context} 1")+
- then have "y = y'" by (rule subst)
+ then have \<open>y = y'\<close> by (rule subst)
} note r' = this
- show "\<forall>y y'. P(y) \<and> P(y') \<longrightarrow> y = y'"
+ show \<open>\<forall>y y'. P(y) \<and> P(y') \<longrightarrow> y = y'\<close>
by (intro strip, elim conjE) (rule r')
qed
qed
-lemma imp_elim: "P \<longrightarrow> Q \<Longrightarrow> (\<not> R \<Longrightarrow> P) \<Longrightarrow> (Q \<Longrightarrow> R) \<Longrightarrow> R"
+lemma imp_elim: \<open>P \<longrightarrow> Q \<Longrightarrow> (\<not> R \<Longrightarrow> P) \<Longrightarrow> (Q \<Longrightarrow> R) \<Longrightarrow> R\<close>
by (rule classical) iprover
-lemma swap: "\<not> P \<Longrightarrow> (\<not> R \<Longrightarrow> P) \<Longrightarrow> R"
+lemma swap: \<open>\<not> P \<Longrightarrow> (\<not> R \<Longrightarrow> P) \<Longrightarrow> R\<close>
by (rule classical) iprover
@@ -214,32 +214,32 @@
\<close>
-lemma ex1_functional: "\<lbrakk>\<exists>! z. P(a,z); P(a,b); P(a,c)\<rbrakk> \<Longrightarrow> b = c"
+lemma ex1_functional: \<open>\<lbrakk>\<exists>! z. P(a,z); P(a,b); P(a,c)\<rbrakk> \<Longrightarrow> b = c\<close>
by blast
text \<open>Elimination of \<open>True\<close> from assumptions:\<close>
-lemma True_implies_equals: "(True \<Longrightarrow> PROP P) \<equiv> PROP P"
+lemma True_implies_equals: \<open>(True \<Longrightarrow> PROP P) \<equiv> PROP P\<close>
proof
- assume "True \<Longrightarrow> PROP P"
- from this and TrueI show "PROP P" .
+ assume \<open>True \<Longrightarrow> PROP P\<close>
+ from this and TrueI show \<open>PROP P\<close> .
next
- assume "PROP P"
- then show "PROP P" .
+ assume \<open>PROP P\<close>
+ then show \<open>PROP P\<close> .
qed
-lemma uncurry: "P \<longrightarrow> Q \<longrightarrow> R \<Longrightarrow> P \<and> Q \<longrightarrow> R"
+lemma uncurry: \<open>P \<longrightarrow> Q \<longrightarrow> R \<Longrightarrow> P \<and> Q \<longrightarrow> R\<close>
by blast
-lemma iff_allI: "(\<And>x. P(x) \<longleftrightarrow> Q(x)) \<Longrightarrow> (\<forall>x. P(x)) \<longleftrightarrow> (\<forall>x. Q(x))"
+lemma iff_allI: \<open>(\<And>x. P(x) \<longleftrightarrow> Q(x)) \<Longrightarrow> (\<forall>x. P(x)) \<longleftrightarrow> (\<forall>x. Q(x))\<close>
by blast
-lemma iff_exI: "(\<And>x. P(x) \<longleftrightarrow> Q(x)) \<Longrightarrow> (\<exists>x. P(x)) \<longleftrightarrow> (\<exists>x. Q(x))"
+lemma iff_exI: \<open>(\<And>x. P(x) \<longleftrightarrow> Q(x)) \<Longrightarrow> (\<exists>x. P(x)) \<longleftrightarrow> (\<exists>x. Q(x))\<close>
by blast
-lemma all_comm: "(\<forall>x y. P(x,y)) \<longleftrightarrow> (\<forall>y x. P(x,y))"
+lemma all_comm: \<open>(\<forall>x y. P(x,y)) \<longleftrightarrow> (\<forall>y x. P(x,y))\<close>
by blast
-lemma ex_comm: "(\<exists>x y. P(x,y)) \<longleftrightarrow> (\<exists>y x. P(x,y))"
+lemma ex_comm: \<open>(\<exists>x y. P(x,y)) \<longleftrightarrow> (\<exists>y x. P(x,y))\<close>
by blast
@@ -251,7 +251,7 @@
false cases boil down to the same thing.
\<close>
-lemma cases_simp: "(P \<longrightarrow> Q) \<and> (\<not> P \<longrightarrow> Q) \<longleftrightarrow> Q"
+lemma cases_simp: \<open>(P \<longrightarrow> Q) \<and> (\<not> P \<longrightarrow> Q) \<longleftrightarrow> Q\<close>
by blast
@@ -267,32 +267,32 @@
text \<open>Existential miniscoping.\<close>
lemma int_ex_simps:
- "\<And>P Q. (\<exists>x. P(x) \<and> Q) \<longleftrightarrow> (\<exists>x. P(x)) \<and> Q"
- "\<And>P Q. (\<exists>x. P \<and> Q(x)) \<longleftrightarrow> P \<and> (\<exists>x. Q(x))"
- "\<And>P Q. (\<exists>x. P(x) \<or> Q) \<longleftrightarrow> (\<exists>x. P(x)) \<or> Q"
- "\<And>P Q. (\<exists>x. P \<or> Q(x)) \<longleftrightarrow> P \<or> (\<exists>x. Q(x))"
+ \<open>\<And>P Q. (\<exists>x. P(x) \<and> Q) \<longleftrightarrow> (\<exists>x. P(x)) \<and> Q\<close>
+ \<open>\<And>P Q. (\<exists>x. P \<and> Q(x)) \<longleftrightarrow> P \<and> (\<exists>x. Q(x))\<close>
+ \<open>\<And>P Q. (\<exists>x. P(x) \<or> Q) \<longleftrightarrow> (\<exists>x. P(x)) \<or> Q\<close>
+ \<open>\<And>P Q. (\<exists>x. P \<or> Q(x)) \<longleftrightarrow> P \<or> (\<exists>x. Q(x))\<close>
by iprover+
text \<open>Classical rules.\<close>
lemma cla_ex_simps:
- "\<And>P Q. (\<exists>x. P(x) \<longrightarrow> Q) \<longleftrightarrow> (\<forall>x. P(x)) \<longrightarrow> Q"
- "\<And>P Q. (\<exists>x. P \<longrightarrow> Q(x)) \<longleftrightarrow> P \<longrightarrow> (\<exists>x. Q(x))"
+ \<open>\<And>P Q. (\<exists>x. P(x) \<longrightarrow> Q) \<longleftrightarrow> (\<forall>x. P(x)) \<longrightarrow> Q\<close>
+ \<open>\<And>P Q. (\<exists>x. P \<longrightarrow> Q(x)) \<longleftrightarrow> P \<longrightarrow> (\<exists>x. Q(x))\<close>
by blast+
lemmas ex_simps = int_ex_simps cla_ex_simps
text \<open>Universal miniscoping.\<close>
lemma int_all_simps:
- "\<And>P Q. (\<forall>x. P(x) \<and> Q) \<longleftrightarrow> (\<forall>x. P(x)) \<and> Q"
- "\<And>P Q. (\<forall>x. P \<and> Q(x)) \<longleftrightarrow> P \<and> (\<forall>x. Q(x))"
- "\<And>P Q. (\<forall>x. P(x) \<longrightarrow> Q) \<longleftrightarrow> (\<exists> x. P(x)) \<longrightarrow> Q"
- "\<And>P Q. (\<forall>x. P \<longrightarrow> Q(x)) \<longleftrightarrow> P \<longrightarrow> (\<forall>x. Q(x))"
+ \<open>\<And>P Q. (\<forall>x. P(x) \<and> Q) \<longleftrightarrow> (\<forall>x. P(x)) \<and> Q\<close>
+ \<open>\<And>P Q. (\<forall>x. P \<and> Q(x)) \<longleftrightarrow> P \<and> (\<forall>x. Q(x))\<close>
+ \<open>\<And>P Q. (\<forall>x. P(x) \<longrightarrow> Q) \<longleftrightarrow> (\<exists> x. P(x)) \<longrightarrow> Q\<close>
+ \<open>\<And>P Q. (\<forall>x. P \<longrightarrow> Q(x)) \<longleftrightarrow> P \<longrightarrow> (\<forall>x. Q(x))\<close>
by iprover+
text \<open>Classical rules.\<close>
lemma cla_all_simps:
- "\<And>P Q. (\<forall>x. P(x) \<or> Q) \<longleftrightarrow> (\<forall>x. P(x)) \<or> Q"
- "\<And>P Q. (\<forall>x. P \<or> Q(x)) \<longleftrightarrow> P \<or> (\<forall>x. Q(x))"
+ \<open>\<And>P Q. (\<forall>x. P(x) \<or> Q) \<longleftrightarrow> (\<forall>x. P(x)) \<or> Q\<close>
+ \<open>\<And>P Q. (\<forall>x. P \<or> Q(x)) \<longleftrightarrow> P \<or> (\<forall>x. Q(x))\<close>
by blast+
lemmas all_simps = int_all_simps cla_all_simps
@@ -300,16 +300,16 @@
subsubsection \<open>Named rewrite rules proved for IFOL\<close>
-lemma imp_disj1: "(P \<longrightarrow> Q) \<or> R \<longleftrightarrow> (P \<longrightarrow> Q \<or> R)" by blast
-lemma imp_disj2: "Q \<or> (P \<longrightarrow> R) \<longleftrightarrow> (P \<longrightarrow> Q \<or> R)" by blast
+lemma imp_disj1: \<open>(P \<longrightarrow> Q) \<or> R \<longleftrightarrow> (P \<longrightarrow> Q \<or> R)\<close> by blast
+lemma imp_disj2: \<open>Q \<or> (P \<longrightarrow> R) \<longleftrightarrow> (P \<longrightarrow> Q \<or> R)\<close> by blast
-lemma de_Morgan_conj: "(\<not> (P \<and> Q)) \<longleftrightarrow> (\<not> P \<or> \<not> Q)" by blast
+lemma de_Morgan_conj: \<open>(\<not> (P \<and> Q)) \<longleftrightarrow> (\<not> P \<or> \<not> Q)\<close> by blast
-lemma not_imp: "\<not> (P \<longrightarrow> Q) \<longleftrightarrow> (P \<and> \<not> Q)" by blast
-lemma not_iff: "\<not> (P \<longleftrightarrow> Q) \<longleftrightarrow> (P \<longleftrightarrow> \<not> Q)" by blast
+lemma not_imp: \<open>\<not> (P \<longrightarrow> Q) \<longleftrightarrow> (P \<and> \<not> Q)\<close> by blast
+lemma not_iff: \<open>\<not> (P \<longleftrightarrow> Q) \<longleftrightarrow> (P \<longleftrightarrow> \<not> Q)\<close> by blast
-lemma not_all: "(\<not> (\<forall>x. P(x))) \<longleftrightarrow> (\<exists>x. \<not> P(x))" by blast
-lemma imp_all: "((\<forall>x. P(x)) \<longrightarrow> Q) \<longleftrightarrow> (\<exists>x. P(x) \<longrightarrow> Q)" by blast
+lemma not_all: \<open>(\<not> (\<forall>x. P(x))) \<longleftrightarrow> (\<exists>x. \<not> P(x))\<close> by blast
+lemma imp_all: \<open>((\<forall>x. P(x)) \<longrightarrow> Q) \<longleftrightarrow> (\<exists>x. P(x) \<longrightarrow> Q)\<close> by blast
lemmas meta_simps =
@@ -320,15 +320,15 @@
refl [THEN P_iff_T] conj_simps disj_simps not_simps
imp_simps iff_simps quant_simps
-lemma notFalseI: "\<not> False" by iprover
+lemma notFalseI: \<open>\<not> False\<close> by iprover
lemma cla_simps_misc:
- "\<not> (P \<and> Q) \<longleftrightarrow> \<not> P \<or> \<not> Q"
- "P \<or> \<not> P"
- "\<not> P \<or> P"
- "\<not> \<not> P \<longleftrightarrow> P"
- "(\<not> P \<longrightarrow> P) \<longleftrightarrow> P"
- "(\<not> P \<longleftrightarrow> \<not> Q) \<longleftrightarrow> (P \<longleftrightarrow> Q)" by blast+
+ \<open>\<not> (P \<and> Q) \<longleftrightarrow> \<not> P \<or> \<not> Q\<close>
+ \<open>P \<or> \<not> P\<close>
+ \<open>\<not> P \<or> P\<close>
+ \<open>\<not> \<not> P \<longleftrightarrow> P\<close>
+ \<open>(\<not> P \<longrightarrow> P) \<longleftrightarrow> P\<close>
+ \<open>(\<not> P \<longleftrightarrow> \<not> Q) \<longleftrightarrow> (P \<longleftrightarrow> Q)\<close> by blast+
lemmas cla_simps =
de_Morgan_conj de_Morgan_disj imp_disj1 imp_disj2
@@ -337,8 +337,8 @@
ML_file "simpdata.ML"
-simproc_setup defined_Ex ("\<exists>x. P(x)") = \<open>fn _ => Quantifier1.rearrange_ex\<close>
-simproc_setup defined_All ("\<forall>x. P(x)") = \<open>fn _ => Quantifier1.rearrange_all\<close>
+simproc_setup defined_Ex (\<open>\<exists>x. P(x)\<close>) = \<open>fn _ => Quantifier1.rearrange_ex\<close>
+simproc_setup defined_All (\<open>\<forall>x. P(x)\<close>) = \<open>fn _ => Quantifier1.rearrange_all\<close>
ML \<open>
(*intuitionistic simprules only*)
@@ -366,35 +366,35 @@
subsection \<open>Other simple lemmas\<close>
-lemma [simp]: "((P \<longrightarrow> R) \<longleftrightarrow> (Q \<longrightarrow> R)) \<longleftrightarrow> ((P \<longleftrightarrow> Q) \<or> R)"
+lemma [simp]: \<open>((P \<longrightarrow> R) \<longleftrightarrow> (Q \<longrightarrow> R)) \<longleftrightarrow> ((P \<longleftrightarrow> Q) \<or> R)\<close>
by blast
-lemma [simp]: "((P \<longrightarrow> Q) \<longleftrightarrow> (P \<longrightarrow> R)) \<longleftrightarrow> (P \<longrightarrow> (Q \<longleftrightarrow> R))"
+lemma [simp]: \<open>((P \<longrightarrow> Q) \<longleftrightarrow> (P \<longrightarrow> R)) \<longleftrightarrow> (P \<longrightarrow> (Q \<longleftrightarrow> R))\<close>
by blast
-lemma not_disj_iff_imp: "\<not> P \<or> Q \<longleftrightarrow> (P \<longrightarrow> Q)"
+lemma not_disj_iff_imp: \<open>\<not> P \<or> Q \<longleftrightarrow> (P \<longrightarrow> Q)\<close>
by blast
subsubsection \<open>Monotonicity of implications\<close>
-lemma conj_mono: "\<lbrakk>P1 \<longrightarrow> Q1; P2 \<longrightarrow> Q2\<rbrakk> \<Longrightarrow> (P1 \<and> P2) \<longrightarrow> (Q1 \<and> Q2)"
+lemma conj_mono: \<open>\<lbrakk>P1 \<longrightarrow> Q1; P2 \<longrightarrow> Q2\<rbrakk> \<Longrightarrow> (P1 \<and> P2) \<longrightarrow> (Q1 \<and> Q2)\<close>
by fast (*or (IntPr.fast_tac 1)*)
-lemma disj_mono: "\<lbrakk>P1 \<longrightarrow> Q1; P2 \<longrightarrow> Q2\<rbrakk> \<Longrightarrow> (P1 \<or> P2) \<longrightarrow> (Q1 \<or> Q2)"
+lemma disj_mono: \<open>\<lbrakk>P1 \<longrightarrow> Q1; P2 \<longrightarrow> Q2\<rbrakk> \<Longrightarrow> (P1 \<or> P2) \<longrightarrow> (Q1 \<or> Q2)\<close>
by fast (*or (IntPr.fast_tac 1)*)
-lemma imp_mono: "\<lbrakk>Q1 \<longrightarrow> P1; P2 \<longrightarrow> Q2\<rbrakk> \<Longrightarrow> (P1 \<longrightarrow> P2) \<longrightarrow> (Q1 \<longrightarrow> Q2)"
+lemma imp_mono: \<open>\<lbrakk>Q1 \<longrightarrow> P1; P2 \<longrightarrow> Q2\<rbrakk> \<Longrightarrow> (P1 \<longrightarrow> P2) \<longrightarrow> (Q1 \<longrightarrow> Q2)\<close>
by fast (*or (IntPr.fast_tac 1)*)
-lemma imp_refl: "P \<longrightarrow> P"
+lemma imp_refl: \<open>P \<longrightarrow> P\<close>
by (rule impI)
text \<open>The quantifier monotonicity rules are also intuitionistically valid.\<close>
-lemma ex_mono: "(\<And>x. P(x) \<longrightarrow> Q(x)) \<Longrightarrow> (\<exists>x. P(x)) \<longrightarrow> (\<exists>x. Q(x))"
+lemma ex_mono: \<open>(\<And>x. P(x) \<longrightarrow> Q(x)) \<Longrightarrow> (\<exists>x. P(x)) \<longrightarrow> (\<exists>x. Q(x))\<close>
by blast
-lemma all_mono: "(\<And>x. P(x) \<longrightarrow> Q(x)) \<Longrightarrow> (\<forall>x. P(x)) \<longrightarrow> (\<forall>x. Q(x))"
+lemma all_mono: \<open>(\<And>x. P(x) \<longrightarrow> Q(x)) \<Longrightarrow> (\<forall>x. P(x)) \<longrightarrow> (\<forall>x. Q(x))\<close>
by blast
@@ -405,21 +405,21 @@
context
begin
-qualified definition "induct_forall(P) \<equiv> \<forall>x. P(x)"
-qualified definition "induct_implies(A, B) \<equiv> A \<longrightarrow> B"
-qualified definition "induct_equal(x, y) \<equiv> x = y"
-qualified definition "induct_conj(A, B) \<equiv> A \<and> B"
+qualified definition \<open>induct_forall(P) \<equiv> \<forall>x. P(x)\<close>
+qualified definition \<open>induct_implies(A, B) \<equiv> A \<longrightarrow> B\<close>
+qualified definition \<open>induct_equal(x, y) \<equiv> x = y\<close>
+qualified definition \<open>induct_conj(A, B) \<equiv> A \<and> B\<close>
-lemma induct_forall_eq: "(\<And>x. P(x)) \<equiv> Trueprop(induct_forall(\<lambda>x. P(x)))"
+lemma induct_forall_eq: \<open>(\<And>x. P(x)) \<equiv> Trueprop(induct_forall(\<lambda>x. P(x)))\<close>
unfolding atomize_all induct_forall_def .
-lemma induct_implies_eq: "(A \<Longrightarrow> B) \<equiv> Trueprop(induct_implies(A, B))"
+lemma induct_implies_eq: \<open>(A \<Longrightarrow> B) \<equiv> Trueprop(induct_implies(A, B))\<close>
unfolding atomize_imp induct_implies_def .
-lemma induct_equal_eq: "(x \<equiv> y) \<equiv> Trueprop(induct_equal(x, y))"
+lemma induct_equal_eq: \<open>(x \<equiv> y) \<equiv> Trueprop(induct_equal(x, y))\<close>
unfolding atomize_eq induct_equal_def .
-lemma induct_conj_eq: "(A &&& B) \<equiv> Trueprop(induct_conj(A, B))"
+lemma induct_conj_eq: \<open>(A &&& B) \<equiv> Trueprop(induct_conj(A, B))\<close>
unfolding atomize_conj induct_conj_def .
lemmas induct_atomize = induct_forall_eq induct_implies_eq induct_equal_eq induct_conj_eq
--- a/src/FOL/IFOL.thy Thu Jan 03 21:48:05 2019 +0100
+++ b/src/FOL/IFOL.thy Thu Jan 03 22:19:19 2019 +0100
@@ -26,76 +26,76 @@
setup Pure_Thy.old_appl_syntax_setup
class "term"
-default_sort "term"
+default_sort \<open>term\<close>
typedecl o
judgment
- Trueprop :: "o \<Rightarrow> prop" (\<open>(_)\<close> 5)
+ Trueprop :: \<open>o \<Rightarrow> prop\<close> (\<open>(_)\<close> 5)
subsubsection \<open>Equality\<close>
axiomatization
- eq :: "['a, 'a] \<Rightarrow> o" (infixl \<open>=\<close> 50)
+ eq :: \<open>['a, 'a] \<Rightarrow> o\<close> (infixl \<open>=\<close> 50)
where
- refl: "a = a" and
- subst: "a = b \<Longrightarrow> P(a) \<Longrightarrow> P(b)"
+ refl: \<open>a = a\<close> and
+ subst: \<open>a = b \<Longrightarrow> P(a) \<Longrightarrow> P(b)\<close>
subsubsection \<open>Propositional logic\<close>
axiomatization
- False :: o and
- conj :: "[o, o] => o" (infixr \<open>\<and>\<close> 35) and
- disj :: "[o, o] => o" (infixr \<open>\<or>\<close> 30) and
- imp :: "[o, o] => o" (infixr \<open>\<longrightarrow>\<close> 25)
+ False :: \<open>o\<close> and
+ conj :: \<open>[o, o] => o\<close> (infixr \<open>\<and>\<close> 35) and
+ disj :: \<open>[o, o] => o\<close> (infixr \<open>\<or>\<close> 30) and
+ imp :: \<open>[o, o] => o\<close> (infixr \<open>\<longrightarrow>\<close> 25)
where
- conjI: "\<lbrakk>P; Q\<rbrakk> \<Longrightarrow> P \<and> Q" and
- conjunct1: "P \<and> Q \<Longrightarrow> P" and
- conjunct2: "P \<and> Q \<Longrightarrow> Q" and
+ conjI: \<open>\<lbrakk>P; Q\<rbrakk> \<Longrightarrow> P \<and> Q\<close> and
+ conjunct1: \<open>P \<and> Q \<Longrightarrow> P\<close> and
+ conjunct2: \<open>P \<and> Q \<Longrightarrow> Q\<close> and
- disjI1: "P \<Longrightarrow> P \<or> Q" and
- disjI2: "Q \<Longrightarrow> P \<or> Q" and
- disjE: "\<lbrakk>P \<or> Q; P \<Longrightarrow> R; Q \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R" and
+ disjI1: \<open>P \<Longrightarrow> P \<or> Q\<close> and
+ disjI2: \<open>Q \<Longrightarrow> P \<or> Q\<close> and
+ disjE: \<open>\<lbrakk>P \<or> Q; P \<Longrightarrow> R; Q \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R\<close> and
- impI: "(P \<Longrightarrow> Q) \<Longrightarrow> P \<longrightarrow> Q" and
- mp: "\<lbrakk>P \<longrightarrow> Q; P\<rbrakk> \<Longrightarrow> Q" and
+ impI: \<open>(P \<Longrightarrow> Q) \<Longrightarrow> P \<longrightarrow> Q\<close> and
+ mp: \<open>\<lbrakk>P \<longrightarrow> Q; P\<rbrakk> \<Longrightarrow> Q\<close> and
- FalseE: "False \<Longrightarrow> P"
+ FalseE: \<open>False \<Longrightarrow> P\<close>
subsubsection \<open>Quantifiers\<close>
axiomatization
- All :: "('a \<Rightarrow> o) \<Rightarrow> o" (binder \<open>\<forall>\<close> 10) and
- Ex :: "('a \<Rightarrow> o) \<Rightarrow> o" (binder \<open>\<exists>\<close> 10)
+ All :: \<open>('a \<Rightarrow> o) \<Rightarrow> o\<close> (binder \<open>\<forall>\<close> 10) and
+ Ex :: \<open>('a \<Rightarrow> o) \<Rightarrow> o\<close> (binder \<open>\<exists>\<close> 10)
where
- allI: "(\<And>x. P(x)) \<Longrightarrow> (\<forall>x. P(x))" and
- spec: "(\<forall>x. P(x)) \<Longrightarrow> P(x)" and
- exI: "P(x) \<Longrightarrow> (\<exists>x. P(x))" and
- exE: "\<lbrakk>\<exists>x. P(x); \<And>x. P(x) \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
+ allI: \<open>(\<And>x. P(x)) \<Longrightarrow> (\<forall>x. P(x))\<close> and
+ spec: \<open>(\<forall>x. P(x)) \<Longrightarrow> P(x)\<close> and
+ exI: \<open>P(x) \<Longrightarrow> (\<exists>x. P(x))\<close> and
+ exE: \<open>\<lbrakk>\<exists>x. P(x); \<And>x. P(x) \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R\<close>
subsubsection \<open>Definitions\<close>
-definition "True \<equiv> False \<longrightarrow> False"
+definition \<open>True \<equiv> False \<longrightarrow> False\<close>
definition Not (\<open>\<not> _\<close> [40] 40)
- where not_def: "\<not> P \<equiv> P \<longrightarrow> False"
+ where not_def: \<open>\<not> P \<equiv> P \<longrightarrow> False\<close>
definition iff (infixr \<open>\<longleftrightarrow>\<close> 25)
- where "P \<longleftrightarrow> Q \<equiv> (P \<longrightarrow> Q) \<and> (Q \<longrightarrow> P)"
+ where \<open>P \<longleftrightarrow> Q \<equiv> (P \<longrightarrow> Q) \<and> (Q \<longrightarrow> P)\<close>
-definition Ex1 :: "('a \<Rightarrow> o) \<Rightarrow> o" (binder \<open>\<exists>!\<close> 10)
- where ex1_def: "\<exists>!x. P(x) \<equiv> \<exists>x. P(x) \<and> (\<forall>y. P(y) \<longrightarrow> y = x)"
+definition Ex1 :: \<open>('a \<Rightarrow> o) \<Rightarrow> o\<close> (binder \<open>\<exists>!\<close> 10)
+ where ex1_def: \<open>\<exists>!x. P(x) \<equiv> \<exists>x. P(x) \<and> (\<forall>y. P(y) \<longrightarrow> y = x)\<close>
axiomatization where \<comment> \<open>Reflection, admissible\<close>
- eq_reflection: "(x = y) \<Longrightarrow> (x \<equiv> y)" and
- iff_reflection: "(P \<longleftrightarrow> Q) \<Longrightarrow> (P \<equiv> Q)"
+ eq_reflection: \<open>(x = y) \<Longrightarrow> (x \<equiv> y)\<close> and
+ iff_reflection: \<open>(P \<longleftrightarrow> Q) \<Longrightarrow> (P \<equiv> Q)\<close>
-abbreviation not_equal :: "['a, 'a] \<Rightarrow> o" (infixl \<open>\<noteq>\<close> 50)
- where "x \<noteq> y \<equiv> \<not> (x = y)"
+abbreviation not_equal :: \<open>['a, 'a] \<Rightarrow> o\<close> (infixl \<open>\<noteq>\<close> 50)
+ where \<open>x \<noteq> y \<equiv> \<not> (x = y)\<close>
subsubsection \<open>Old-style ASCII syntax\<close>
@@ -116,44 +116,44 @@
lemmas strip = impI allI
-lemma TrueI: True
+lemma TrueI: \<open>True\<close>
unfolding True_def by (rule impI)
subsubsection \<open>Sequent-style elimination rules for \<open>\<and>\<close> \<open>\<longrightarrow>\<close> and \<open>\<forall>\<close>\<close>
lemma conjE:
- assumes major: "P \<and> Q"
- and r: "\<lbrakk>P; Q\<rbrakk> \<Longrightarrow> R"
- shows R
+ assumes major: \<open>P \<and> Q\<close>
+ and r: \<open>\<lbrakk>P; Q\<rbrakk> \<Longrightarrow> R\<close>
+ shows \<open>R\<close>
apply (rule r)
apply (rule major [THEN conjunct1])
apply (rule major [THEN conjunct2])
done
lemma impE:
- assumes major: "P \<longrightarrow> Q"
- and P
- and r: "Q \<Longrightarrow> R"
- shows R
+ assumes major: \<open>P \<longrightarrow> Q\<close>
+ and \<open>P\<close>
+ and r: \<open>Q \<Longrightarrow> R\<close>
+ shows \<open>R\<close>
apply (rule r)
apply (rule major [THEN mp])
apply (rule \<open>P\<close>)
done
lemma allE:
- assumes major: "\<forall>x. P(x)"
- and r: "P(x) \<Longrightarrow> R"
- shows R
+ assumes major: \<open>\<forall>x. P(x)\<close>
+ and r: \<open>P(x) \<Longrightarrow> R\<close>
+ shows \<open>R\<close>
apply (rule r)
apply (rule major [THEN spec])
done
text \<open>Duplicates the quantifier; for use with @{ML eresolve_tac}.\<close>
lemma all_dupE:
- assumes major: "\<forall>x. P(x)"
- and r: "\<lbrakk>P(x); \<forall>x. P(x)\<rbrakk> \<Longrightarrow> R"
- shows R
+ assumes major: \<open>\<forall>x. P(x)\<close>
+ and r: \<open>\<lbrakk>P(x); \<forall>x. P(x)\<rbrakk> \<Longrightarrow> R\<close>
+ shows \<open>R\<close>
apply (rule r)
apply (rule major [THEN spec])
apply (rule major)
@@ -162,20 +162,20 @@
subsubsection \<open>Negation rules, which translate between \<open>\<not> P\<close> and \<open>P \<longrightarrow> False\<close>\<close>
-lemma notI: "(P \<Longrightarrow> False) \<Longrightarrow> \<not> P"
+lemma notI: \<open>(P \<Longrightarrow> False) \<Longrightarrow> \<not> P\<close>
unfolding not_def by (erule impI)
-lemma notE: "\<lbrakk>\<not> P; P\<rbrakk> \<Longrightarrow> R"
+lemma notE: \<open>\<lbrakk>\<not> P; P\<rbrakk> \<Longrightarrow> R\<close>
unfolding not_def by (erule mp [THEN FalseE])
-lemma rev_notE: "\<lbrakk>P; \<not> P\<rbrakk> \<Longrightarrow> R"
+lemma rev_notE: \<open>\<lbrakk>P; \<not> P\<rbrakk> \<Longrightarrow> R\<close>
by (erule notE)
text \<open>This is useful with the special implication rules for each kind of \<open>P\<close>.\<close>
lemma not_to_imp:
- assumes "\<not> P"
- and r: "P \<longrightarrow> False \<Longrightarrow> Q"
- shows Q
+ assumes \<open>\<not> P\<close>
+ and r: \<open>P \<longrightarrow> False \<Longrightarrow> Q\<close>
+ shows \<open>Q\<close>
apply (rule r)
apply (rule impI)
apply (erule notE [OF \<open>\<not> P\<close>])
@@ -185,14 +185,14 @@
For substitution into an assumption \<open>P\<close>, reduce \<open>Q\<close> to \<open>P \<longrightarrow> Q\<close>, substitute into this implication, then apply \<open>impI\<close> to
move \<open>P\<close> back into the assumptions.
\<close>
-lemma rev_mp: "\<lbrakk>P; P \<longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
+lemma rev_mp: \<open>\<lbrakk>P; P \<longrightarrow> Q\<rbrakk> \<Longrightarrow> Q\<close>
by (erule mp)
text \<open>Contrapositive of an inference rule.\<close>
lemma contrapos:
- assumes major: "\<not> Q"
- and minor: "P \<Longrightarrow> Q"
- shows "\<not> P"
+ assumes major: \<open>\<not> Q\<close>
+ and minor: \<open>P \<Longrightarrow> Q\<close>
+ shows \<open>\<not> P\<close>
apply (rule major [THEN notE, THEN notI])
apply (erule minor)
done
@@ -214,7 +214,7 @@
subsection \<open>If-and-only-if\<close>
-lemma iffI: "\<lbrakk>P \<Longrightarrow> Q; Q \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P \<longleftrightarrow> Q"
+lemma iffI: \<open>\<lbrakk>P \<Longrightarrow> Q; Q \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P \<longleftrightarrow> Q\<close>
apply (unfold iff_def)
apply (rule conjI)
apply (erule impI)
@@ -222,9 +222,9 @@
done
lemma iffE:
- assumes major: "P \<longleftrightarrow> Q"
- and r: "P \<longrightarrow> Q \<Longrightarrow> Q \<longrightarrow> P \<Longrightarrow> R"
- shows R
+ assumes major: \<open>P \<longleftrightarrow> Q\<close>
+ and r: \<open>P \<longrightarrow> Q \<Longrightarrow> Q \<longrightarrow> P \<Longrightarrow> R\<close>
+ shows \<open>R\<close>
apply (insert major, unfold iff_def)
apply (erule conjE)
apply (erule r)
@@ -234,38 +234,38 @@
subsubsection \<open>Destruct rules for \<open>\<longleftrightarrow>\<close> similar to Modus Ponens\<close>
-lemma iffD1: "\<lbrakk>P \<longleftrightarrow> Q; P\<rbrakk> \<Longrightarrow> Q"
+lemma iffD1: \<open>\<lbrakk>P \<longleftrightarrow> Q; P\<rbrakk> \<Longrightarrow> Q\<close>
apply (unfold iff_def)
apply (erule conjunct1 [THEN mp])
apply assumption
done
-lemma iffD2: "\<lbrakk>P \<longleftrightarrow> Q; Q\<rbrakk> \<Longrightarrow> P"
+lemma iffD2: \<open>\<lbrakk>P \<longleftrightarrow> Q; Q\<rbrakk> \<Longrightarrow> P\<close>
apply (unfold iff_def)
apply (erule conjunct2 [THEN mp])
apply assumption
done
-lemma rev_iffD1: "\<lbrakk>P; P \<longleftrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
+lemma rev_iffD1: \<open>\<lbrakk>P; P \<longleftrightarrow> Q\<rbrakk> \<Longrightarrow> Q\<close>
apply (erule iffD1)
apply assumption
done
-lemma rev_iffD2: "\<lbrakk>Q; P \<longleftrightarrow> Q\<rbrakk> \<Longrightarrow> P"
+lemma rev_iffD2: \<open>\<lbrakk>Q; P \<longleftrightarrow> Q\<rbrakk> \<Longrightarrow> P\<close>
apply (erule iffD2)
apply assumption
done
-lemma iff_refl: "P \<longleftrightarrow> P"
+lemma iff_refl: \<open>P \<longleftrightarrow> P\<close>
by (rule iffI)
-lemma iff_sym: "Q \<longleftrightarrow> P \<Longrightarrow> P \<longleftrightarrow> Q"
+lemma iff_sym: \<open>Q \<longleftrightarrow> P \<Longrightarrow> P \<longleftrightarrow> Q\<close>
apply (erule iffE)
apply (rule iffI)
apply (assumption | erule mp)+
done
-lemma iff_trans: "\<lbrakk>P \<longleftrightarrow> Q; Q \<longleftrightarrow> R\<rbrakk> \<Longrightarrow> P \<longleftrightarrow> R"
+lemma iff_trans: \<open>\<lbrakk>P \<longleftrightarrow> Q; Q \<longleftrightarrow> R\<rbrakk> \<Longrightarrow> P \<longleftrightarrow> R\<close>
apply (rule iffI)
apply (assumption | erule iffE | erule (1) notE impE)+
done
@@ -282,20 +282,20 @@
do NOT mean the same thing. The parser treats \<open>\<exists>!x y.P(x,y)\<close> as sequential.
\<close>
-lemma ex1I: "P(a) \<Longrightarrow> (\<And>x. P(x) \<Longrightarrow> x = a) \<Longrightarrow> \<exists>!x. P(x)"
+lemma ex1I: \<open>P(a) \<Longrightarrow> (\<And>x. P(x) \<Longrightarrow> x = a) \<Longrightarrow> \<exists>!x. P(x)\<close>
apply (unfold ex1_def)
apply (assumption | rule exI conjI allI impI)+
done
text \<open>Sometimes easier to use: the premises have no shared variables. Safe!\<close>
-lemma ex_ex1I: "\<exists>x. P(x) \<Longrightarrow> (\<And>x y. \<lbrakk>P(x); P(y)\<rbrakk> \<Longrightarrow> x = y) \<Longrightarrow> \<exists>!x. P(x)"
+lemma ex_ex1I: \<open>\<exists>x. P(x) \<Longrightarrow> (\<And>x y. \<lbrakk>P(x); P(y)\<rbrakk> \<Longrightarrow> x = y) \<Longrightarrow> \<exists>!x. P(x)\<close>
apply (erule exE)
apply (rule ex1I)
apply assumption
apply assumption
done
-lemma ex1E: "\<exists>! x. P(x) \<Longrightarrow> (\<And>x. \<lbrakk>P(x); \<forall>y. P(y) \<longrightarrow> y = x\<rbrakk> \<Longrightarrow> R) \<Longrightarrow> R"
+lemma ex1E: \<open>\<exists>! x. P(x) \<Longrightarrow> (\<And>x. \<lbrakk>P(x); \<forall>y. P(y) \<longrightarrow> y = x\<rbrakk> \<Longrightarrow> R) \<Longrightarrow> R\<close>
apply (unfold ex1_def)
apply (assumption | erule exE conjE)+
done
@@ -315,77 +315,77 @@
(fn prems => fn ctxt => SIMPLE_METHOD' (iff_tac ctxt prems))\<close>
lemma conj_cong:
- assumes "P \<longleftrightarrow> P'"
- and "P' \<Longrightarrow> Q \<longleftrightarrow> Q'"
- shows "(P \<and> Q) \<longleftrightarrow> (P' \<and> Q')"
+ assumes \<open>P \<longleftrightarrow> P'\<close>
+ and \<open>P' \<Longrightarrow> Q \<longleftrightarrow> Q'\<close>
+ shows \<open>(P \<and> Q) \<longleftrightarrow> (P' \<and> Q')\<close>
apply (insert assms)
apply (assumption | rule iffI conjI | erule iffE conjE mp | iff assms)+
done
text \<open>Reversed congruence rule! Used in ZF/Order.\<close>
lemma conj_cong2:
- assumes "P \<longleftrightarrow> P'"
- and "P' \<Longrightarrow> Q \<longleftrightarrow> Q'"
- shows "(Q \<and> P) \<longleftrightarrow> (Q' \<and> P')"
+ assumes \<open>P \<longleftrightarrow> P'\<close>
+ and \<open>P' \<Longrightarrow> Q \<longleftrightarrow> Q'\<close>
+ shows \<open>(Q \<and> P) \<longleftrightarrow> (Q' \<and> P')\<close>
apply (insert assms)
apply (assumption | rule iffI conjI | erule iffE conjE mp | iff assms)+
done
lemma disj_cong:
- assumes "P \<longleftrightarrow> P'" and "Q \<longleftrightarrow> Q'"
- shows "(P \<or> Q) \<longleftrightarrow> (P' \<or> Q')"
+ assumes \<open>P \<longleftrightarrow> P'\<close> and \<open>Q \<longleftrightarrow> Q'\<close>
+ shows \<open>(P \<or> Q) \<longleftrightarrow> (P' \<or> Q')\<close>
apply (insert assms)
apply (erule iffE disjE disjI1 disjI2 |
assumption | rule iffI | erule (1) notE impE)+
done
lemma imp_cong:
- assumes "P \<longleftrightarrow> P'"
- and "P' \<Longrightarrow> Q \<longleftrightarrow> Q'"
- shows "(P \<longrightarrow> Q) \<longleftrightarrow> (P' \<longrightarrow> Q')"
+ assumes \<open>P \<longleftrightarrow> P'\<close>
+ and \<open>P' \<Longrightarrow> Q \<longleftrightarrow> Q'\<close>
+ shows \<open>(P \<longrightarrow> Q) \<longleftrightarrow> (P' \<longrightarrow> Q')\<close>
apply (insert assms)
apply (assumption | rule iffI impI | erule iffE | erule (1) notE impE | iff assms)+
done
-lemma iff_cong: "\<lbrakk>P \<longleftrightarrow> P'; Q \<longleftrightarrow> Q'\<rbrakk> \<Longrightarrow> (P \<longleftrightarrow> Q) \<longleftrightarrow> (P' \<longleftrightarrow> Q')"
+lemma iff_cong: \<open>\<lbrakk>P \<longleftrightarrow> P'; Q \<longleftrightarrow> Q'\<rbrakk> \<Longrightarrow> (P \<longleftrightarrow> Q) \<longleftrightarrow> (P' \<longleftrightarrow> Q')\<close>
apply (erule iffE | assumption | rule iffI | erule (1) notE impE)+
done
-lemma not_cong: "P \<longleftrightarrow> P' \<Longrightarrow> \<not> P \<longleftrightarrow> \<not> P'"
+lemma not_cong: \<open>P \<longleftrightarrow> P' \<Longrightarrow> \<not> P \<longleftrightarrow> \<not> P'\<close>
apply (assumption | rule iffI notI | erule (1) notE impE | erule iffE notE)+
done
lemma all_cong:
- assumes "\<And>x. P(x) \<longleftrightarrow> Q(x)"
- shows "(\<forall>x. P(x)) \<longleftrightarrow> (\<forall>x. Q(x))"
+ assumes \<open>\<And>x. P(x) \<longleftrightarrow> Q(x)\<close>
+ shows \<open>(\<forall>x. P(x)) \<longleftrightarrow> (\<forall>x. Q(x))\<close>
apply (assumption | rule iffI allI | erule (1) notE impE | erule allE | iff assms)+
done
lemma ex_cong:
- assumes "\<And>x. P(x) \<longleftrightarrow> Q(x)"
- shows "(\<exists>x. P(x)) \<longleftrightarrow> (\<exists>x. Q(x))"
+ assumes \<open>\<And>x. P(x) \<longleftrightarrow> Q(x)\<close>
+ shows \<open>(\<exists>x. P(x)) \<longleftrightarrow> (\<exists>x. Q(x))\<close>
apply (erule exE | assumption | rule iffI exI | erule (1) notE impE | iff assms)+
done
lemma ex1_cong:
- assumes "\<And>x. P(x) \<longleftrightarrow> Q(x)"
- shows "(\<exists>!x. P(x)) \<longleftrightarrow> (\<exists>!x. Q(x))"
+ assumes \<open>\<And>x. P(x) \<longleftrightarrow> Q(x)\<close>
+ shows \<open>(\<exists>!x. P(x)) \<longleftrightarrow> (\<exists>!x. Q(x))\<close>
apply (erule ex1E spec [THEN mp] | assumption | rule iffI ex1I | erule (1) notE impE | iff assms)+
done
subsection \<open>Equality rules\<close>
-lemma sym: "a = b \<Longrightarrow> b = a"
+lemma sym: \<open>a = b \<Longrightarrow> b = a\<close>
apply (erule subst)
apply (rule refl)
done
-lemma trans: "\<lbrakk>a = b; b = c\<rbrakk> \<Longrightarrow> a = c"
+lemma trans: \<open>\<lbrakk>a = b; b = c\<rbrakk> \<Longrightarrow> a = c\<close>
apply (erule subst, assumption)
done
-lemma not_sym: "b \<noteq> a \<Longrightarrow> a \<noteq> b"
+lemma not_sym: \<open>b \<noteq> a \<Longrightarrow> a \<noteq> b\<close>
apply (erule contrapos)
apply (erule sym)
done
@@ -395,28 +395,28 @@
the first for definitions of formulae and the second for terms.
\<close>
-lemma def_imp_iff: "(A \<equiv> B) \<Longrightarrow> A \<longleftrightarrow> B"
+lemma def_imp_iff: \<open>(A \<equiv> B) \<Longrightarrow> A \<longleftrightarrow> B\<close>
apply unfold
apply (rule iff_refl)
done
-lemma meta_eq_to_obj_eq: "(A \<equiv> B) \<Longrightarrow> A = B"
+lemma meta_eq_to_obj_eq: \<open>(A \<equiv> B) \<Longrightarrow> A = B\<close>
apply unfold
apply (rule refl)
done
-lemma meta_eq_to_iff: "x \<equiv> y \<Longrightarrow> x \<longleftrightarrow> y"
+lemma meta_eq_to_iff: \<open>x \<equiv> y \<Longrightarrow> x \<longleftrightarrow> y\<close>
by unfold (rule iff_refl)
text \<open>Substitution.\<close>
-lemma ssubst: "\<lbrakk>b = a; P(a)\<rbrakk> \<Longrightarrow> P(b)"
+lemma ssubst: \<open>\<lbrakk>b = a; P(a)\<rbrakk> \<Longrightarrow> P(b)\<close>
apply (drule sym)
apply (erule (1) subst)
done
text \<open>A special case of \<open>ex1E\<close> that would otherwise need quantifier
expansion.\<close>
-lemma ex1_equalsE: "\<lbrakk>\<exists>!x. P(x); P(a); P(b)\<rbrakk> \<Longrightarrow> a = b"
+lemma ex1_equalsE: \<open>\<lbrakk>\<exists>!x. P(x); P(a); P(b)\<rbrakk> \<Longrightarrow> a = b\<close>
apply (erule ex1E)
apply (rule trans)
apply (rule_tac [2] sym)
@@ -426,17 +426,17 @@
subsubsection \<open>Polymorphic congruence rules\<close>
-lemma subst_context: "a = b \<Longrightarrow> t(a) = t(b)"
+lemma subst_context: \<open>a = b \<Longrightarrow> t(a) = t(b)\<close>
apply (erule ssubst)
apply (rule refl)
done
-lemma subst_context2: "\<lbrakk>a = b; c = d\<rbrakk> \<Longrightarrow> t(a,c) = t(b,d)"
+lemma subst_context2: \<open>\<lbrakk>a = b; c = d\<rbrakk> \<Longrightarrow> t(a,c) = t(b,d)\<close>
apply (erule ssubst)+
apply (rule refl)
done
-lemma subst_context3: "\<lbrakk>a = b; c = d; e = f\<rbrakk> \<Longrightarrow> t(a,c,e) = t(b,d,f)"
+lemma subst_context3: \<open>\<lbrakk>a = b; c = d; e = f\<rbrakk> \<Longrightarrow> t(a,c,e) = t(b,d,f)\<close>
apply (erule ssubst)+
apply (rule refl)
done
@@ -449,7 +449,7 @@
| |
c = d
\<close>
-lemma box_equals: "\<lbrakk>a = b; a = c; b = d\<rbrakk> \<Longrightarrow> c = d"
+lemma box_equals: \<open>\<lbrakk>a = b; a = c; b = d\<rbrakk> \<Longrightarrow> c = d\<close>
apply (rule trans)
apply (rule trans)
apply (rule sym)
@@ -457,7 +457,7 @@
done
text \<open>Dual of \<open>box_equals\<close>: for proving equalities backwards.\<close>
-lemma simp_equals: "\<lbrakk>a = c; b = d; c = d\<rbrakk> \<Longrightarrow> a = b"
+lemma simp_equals: \<open>\<lbrakk>a = c; b = d; c = d\<rbrakk> \<Longrightarrow> a = b\<close>
apply (rule trans)
apply (rule trans)
apply assumption+
@@ -467,13 +467,13 @@
subsubsection \<open>Congruence rules for predicate letters\<close>
-lemma pred1_cong: "a = a' \<Longrightarrow> P(a) \<longleftrightarrow> P(a')"
+lemma pred1_cong: \<open>a = a' \<Longrightarrow> P(a) \<longleftrightarrow> P(a')\<close>
apply (rule iffI)
apply (erule (1) subst)
apply (erule (1) ssubst)
done
-lemma pred2_cong: "\<lbrakk>a = a'; b = b'\<rbrakk> \<Longrightarrow> P(a,b) \<longleftrightarrow> P(a',b')"
+lemma pred2_cong: \<open>\<lbrakk>a = a'; b = b'\<rbrakk> \<Longrightarrow> P(a,b) \<longleftrightarrow> P(a',b')\<close>
apply (rule iffI)
apply (erule subst)+
apply assumption
@@ -481,7 +481,7 @@
apply assumption
done
-lemma pred3_cong: "\<lbrakk>a = a'; b = b'; c = c'\<rbrakk> \<Longrightarrow> P(a,b,c) \<longleftrightarrow> P(a',b',c')"
+lemma pred3_cong: \<open>\<lbrakk>a = a'; b = b'; c = c'\<rbrakk> \<Longrightarrow> P(a,b,c) \<longleftrightarrow> P(a',b',c')\<close>
apply (rule iffI)
apply (erule subst)+
apply assumption
@@ -490,7 +490,7 @@
done
text \<open>Special case for the equality predicate!\<close>
-lemma eq_cong: "\<lbrakk>a = a'; b = b'\<rbrakk> \<Longrightarrow> a = b \<longleftrightarrow> a' = b'"
+lemma eq_cong: \<open>\<lbrakk>a = a'; b = b'\<rbrakk> \<Longrightarrow> a = b \<longleftrightarrow> a' = b'\<close>
apply (erule (1) pred2_cong)
done
@@ -507,29 +507,29 @@
\<close>
lemma conj_impE:
- assumes major: "(P \<and> Q) \<longrightarrow> S"
- and r: "P \<longrightarrow> (Q \<longrightarrow> S) \<Longrightarrow> R"
- shows R
+ assumes major: \<open>(P \<and> Q) \<longrightarrow> S\<close>
+ and r: \<open>P \<longrightarrow> (Q \<longrightarrow> S) \<Longrightarrow> R\<close>
+ shows \<open>R\<close>
by (assumption | rule conjI impI major [THEN mp] r)+
lemma disj_impE:
- assumes major: "(P \<or> Q) \<longrightarrow> S"
- and r: "\<lbrakk>P \<longrightarrow> S; Q \<longrightarrow> S\<rbrakk> \<Longrightarrow> R"
- shows R
+ assumes major: \<open>(P \<or> Q) \<longrightarrow> S\<close>
+ and r: \<open>\<lbrakk>P \<longrightarrow> S; Q \<longrightarrow> S\<rbrakk> \<Longrightarrow> R\<close>
+ shows \<open>R\<close>
by (assumption | rule disjI1 disjI2 impI major [THEN mp] r)+
text \<open>Simplifies the implication. Classical version is stronger.
Still UNSAFE since Q must be provable -- backtracking needed.\<close>
lemma imp_impE:
- assumes major: "(P \<longrightarrow> Q) \<longrightarrow> S"
- and r1: "\<lbrakk>P; Q \<longrightarrow> S\<rbrakk> \<Longrightarrow> Q"
- and r2: "S \<Longrightarrow> R"
- shows R
+ assumes major: \<open>(P \<longrightarrow> Q) \<longrightarrow> S\<close>
+ and r1: \<open>\<lbrakk>P; Q \<longrightarrow> S\<rbrakk> \<Longrightarrow> Q\<close>
+ and r2: \<open>S \<Longrightarrow> R\<close>
+ shows \<open>R\<close>
by (assumption | rule impI major [THEN mp] r1 r2)+
text \<open>Simplifies the implication. Classical version is stronger.
Still UNSAFE since ~P must be provable -- backtracking needed.\<close>
-lemma not_impE: "\<not> P \<longrightarrow> S \<Longrightarrow> (P \<Longrightarrow> False) \<Longrightarrow> (S \<Longrightarrow> R) \<Longrightarrow> R"
+lemma not_impE: \<open>\<not> P \<longrightarrow> S \<Longrightarrow> (P \<Longrightarrow> False) \<Longrightarrow> (S \<Longrightarrow> R) \<Longrightarrow> R\<close>
apply (drule mp)
apply (rule notI)
apply assumption
@@ -538,21 +538,21 @@
text \<open>Simplifies the implication. UNSAFE.\<close>
lemma iff_impE:
- assumes major: "(P \<longleftrightarrow> Q) \<longrightarrow> S"
- and r1: "\<lbrakk>P; Q \<longrightarrow> S\<rbrakk> \<Longrightarrow> Q"
- and r2: "\<lbrakk>Q; P \<longrightarrow> S\<rbrakk> \<Longrightarrow> P"
- and r3: "S \<Longrightarrow> R"
- shows R
+ assumes major: \<open>(P \<longleftrightarrow> Q) \<longrightarrow> S\<close>
+ and r1: \<open>\<lbrakk>P; Q \<longrightarrow> S\<rbrakk> \<Longrightarrow> Q\<close>
+ and r2: \<open>\<lbrakk>Q; P \<longrightarrow> S\<rbrakk> \<Longrightarrow> P\<close>
+ and r3: \<open>S \<Longrightarrow> R\<close>
+ shows \<open>R\<close>
apply (assumption | rule iffI impI major [THEN mp] r1 r2 r3)+
done
text \<open>What if \<open>(\<forall>x. \<not> \<not> P(x)) \<longrightarrow> \<not> \<not> (\<forall>x. P(x))\<close> is an assumption?
UNSAFE.\<close>
lemma all_impE:
- assumes major: "(\<forall>x. P(x)) \<longrightarrow> S"
- and r1: "\<And>x. P(x)"
- and r2: "S \<Longrightarrow> R"
- shows R
+ assumes major: \<open>(\<forall>x. P(x)) \<longrightarrow> S\<close>
+ and r1: \<open>\<And>x. P(x)\<close>
+ and r2: \<open>S \<Longrightarrow> R\<close>
+ shows \<open>R\<close>
apply (rule allI impI major [THEN mp] r1 r2)+
done
@@ -560,14 +560,14 @@
Unsafe: \<open>\<exists>x. P(x)) \<longrightarrow> S\<close> is equivalent
to \<open>\<forall>x. P(x) \<longrightarrow> S\<close>.\<close>
lemma ex_impE:
- assumes major: "(\<exists>x. P(x)) \<longrightarrow> S"
- and r: "P(x) \<longrightarrow> S \<Longrightarrow> R"
- shows R
+ assumes major: \<open>(\<exists>x. P(x)) \<longrightarrow> S\<close>
+ and r: \<open>P(x) \<longrightarrow> S \<Longrightarrow> R\<close>
+ shows \<open>R\<close>
apply (assumption | rule exI impI major [THEN mp] r)+
done
text \<open>Courtesy of Krzysztof Grabczewski.\<close>
-lemma disj_imp_disj: "P \<or> Q \<Longrightarrow> (P \<Longrightarrow> R) \<Longrightarrow> (Q \<Longrightarrow> S) \<Longrightarrow> R \<or> S"
+lemma disj_imp_disj: \<open>P \<or> Q \<Longrightarrow> (P \<Longrightarrow> R) \<Longrightarrow> (Q \<Longrightarrow> S) \<Longrightarrow> R \<or> S\<close>
apply (erule disjE)
apply (rule disjI1) apply assumption
apply (rule disjI2) apply assumption
@@ -584,7 +584,7 @@
ML_file "fologic.ML"
-lemma thin_refl: "\<lbrakk>x = x; PROP W\<rbrakk> \<Longrightarrow> PROP W" .
+lemma thin_refl: \<open>\<lbrakk>x = x; PROP W\<rbrakk> \<Longrightarrow> PROP W\<close> .
ML \<open>
structure Hypsubst = Hypsubst
@@ -611,32 +611,32 @@
setup \<open>Intuitionistic.method_setup @{binding iprover}\<close>
lemma impE':
- assumes 1: "P \<longrightarrow> Q"
- and 2: "Q \<Longrightarrow> R"
- and 3: "P \<longrightarrow> Q \<Longrightarrow> P"
- shows R
+ assumes 1: \<open>P \<longrightarrow> Q\<close>
+ and 2: \<open>Q \<Longrightarrow> R\<close>
+ and 3: \<open>P \<longrightarrow> Q \<Longrightarrow> P\<close>
+ shows \<open>R\<close>
proof -
- from 3 and 1 have P .
- with 1 have Q by (rule impE)
- with 2 show R .
+ from 3 and 1 have \<open>P\<close> .
+ with 1 have \<open>Q\<close> by (rule impE)
+ with 2 show \<open>R\<close> .
qed
lemma allE':
- assumes 1: "\<forall>x. P(x)"
- and 2: "P(x) \<Longrightarrow> \<forall>x. P(x) \<Longrightarrow> Q"
- shows Q
+ assumes 1: \<open>\<forall>x. P(x)\<close>
+ and 2: \<open>P(x) \<Longrightarrow> \<forall>x. P(x) \<Longrightarrow> Q\<close>
+ shows \<open>Q\<close>
proof -
- from 1 have "P(x)" by (rule spec)
- from this and 1 show Q by (rule 2)
+ from 1 have \<open>P(x)\<close> by (rule spec)
+ from this and 1 show \<open>Q\<close> by (rule 2)
qed
lemma notE':
- assumes 1: "\<not> P"
- and 2: "\<not> P \<Longrightarrow> P"
- shows R
+ assumes 1: \<open>\<not> P\<close>
+ and 2: \<open>\<not> P \<Longrightarrow> P\<close>
+ shows \<open>R\<close>
proof -
- from 2 and 1 have P .
- with 1 show R by (rule notE)
+ from 2 and 1 have \<open>P\<close> .
+ with 1 show \<open>R\<close> by (rule notE)
qed
lemmas [Pure.elim!] = disjE iffE FalseE conjE exE
@@ -650,14 +650,14 @@
\<close>
-lemma iff_not_sym: "\<not> (Q \<longleftrightarrow> P) \<Longrightarrow> \<not> (P \<longleftrightarrow> Q)"
+lemma iff_not_sym: \<open>\<not> (Q \<longleftrightarrow> P) \<Longrightarrow> \<not> (P \<longleftrightarrow> Q)\<close>
by iprover
lemmas [sym] = sym iff_sym not_sym iff_not_sym
and [Pure.elim?] = iffD1 iffD2 impE
-lemma eq_commute: "a = b \<longleftrightarrow> b = a"
+lemma eq_commute: \<open>a = b \<longleftrightarrow> b = a\<close>
apply (rule iffI)
apply (erule sym)+
done
@@ -665,56 +665,56 @@
subsection \<open>Atomizing meta-level rules\<close>
-lemma atomize_all [atomize]: "(\<And>x. P(x)) \<equiv> Trueprop (\<forall>x. P(x))"
+lemma atomize_all [atomize]: \<open>(\<And>x. P(x)) \<equiv> Trueprop (\<forall>x. P(x))\<close>
proof
- assume "\<And>x. P(x)"
- then show "\<forall>x. P(x)" ..
+ assume \<open>\<And>x. P(x)\<close>
+ then show \<open>\<forall>x. P(x)\<close> ..
next
- assume "\<forall>x. P(x)"
- then show "\<And>x. P(x)" ..
+ assume \<open>\<forall>x. P(x)\<close>
+ then show \<open>\<And>x. P(x)\<close> ..
qed
-lemma atomize_imp [atomize]: "(A \<Longrightarrow> B) \<equiv> Trueprop (A \<longrightarrow> B)"
+lemma atomize_imp [atomize]: \<open>(A \<Longrightarrow> B) \<equiv> Trueprop (A \<longrightarrow> B)\<close>
proof
- assume "A \<Longrightarrow> B"
- then show "A \<longrightarrow> B" ..
+ assume \<open>A \<Longrightarrow> B\<close>
+ then show \<open>A \<longrightarrow> B\<close> ..
next
- assume "A \<longrightarrow> B" and A
- then show B by (rule mp)
+ assume \<open>A \<longrightarrow> B\<close> and \<open>A\<close>
+ then show \<open>B\<close> by (rule mp)
qed
-lemma atomize_eq [atomize]: "(x \<equiv> y) \<equiv> Trueprop (x = y)"
+lemma atomize_eq [atomize]: \<open>(x \<equiv> y) \<equiv> Trueprop (x = y)\<close>
proof
- assume "x \<equiv> y"
- show "x = y" unfolding \<open>x \<equiv> y\<close> by (rule refl)
+ assume \<open>x \<equiv> y\<close>
+ show \<open>x = y\<close> unfolding \<open>x \<equiv> y\<close> by (rule refl)
next
- assume "x = y"
- then show "x \<equiv> y" by (rule eq_reflection)
+ assume \<open>x = y\<close>
+ then show \<open>x \<equiv> y\<close> by (rule eq_reflection)
qed
-lemma atomize_iff [atomize]: "(A \<equiv> B) \<equiv> Trueprop (A \<longleftrightarrow> B)"
+lemma atomize_iff [atomize]: \<open>(A \<equiv> B) \<equiv> Trueprop (A \<longleftrightarrow> B)\<close>
proof
- assume "A \<equiv> B"
- show "A \<longleftrightarrow> B" unfolding \<open>A \<equiv> B\<close> by (rule iff_refl)
+ assume \<open>A \<equiv> B\<close>
+ show \<open>A \<longleftrightarrow> B\<close> unfolding \<open>A \<equiv> B\<close> by (rule iff_refl)
next
- assume "A \<longleftrightarrow> B"
- then show "A \<equiv> B" by (rule iff_reflection)
+ assume \<open>A \<longleftrightarrow> B\<close>
+ then show \<open>A \<equiv> B\<close> by (rule iff_reflection)
qed
-lemma atomize_conj [atomize]: "(A &&& B) \<equiv> Trueprop (A \<and> B)"
+lemma atomize_conj [atomize]: \<open>(A &&& B) \<equiv> Trueprop (A \<and> B)\<close>
proof
- assume conj: "A &&& B"
- show "A \<and> B"
+ assume conj: \<open>A &&& B\<close>
+ show \<open>A \<and> B\<close>
proof (rule conjI)
- from conj show A by (rule conjunctionD1)
- from conj show B by (rule conjunctionD2)
+ from conj show \<open>A\<close> by (rule conjunctionD1)
+ from conj show \<open>B\<close> by (rule conjunctionD2)
qed
next
- assume conj: "A \<and> B"
- show "A &&& B"
+ assume conj: \<open>A \<and> B\<close>
+ show \<open>A &&& B\<close>
proof -
- from conj show A ..
- from conj show B ..
+ from conj show \<open>A\<close> ..
+ from conj show \<open>B\<close> ..
qed
qed
@@ -724,24 +724,24 @@
subsection \<open>Atomizing elimination rules\<close>
-lemma atomize_exL[atomize_elim]: "(\<And>x. P(x) \<Longrightarrow> Q) \<equiv> ((\<exists>x. P(x)) \<Longrightarrow> Q)"
+lemma atomize_exL[atomize_elim]: \<open>(\<And>x. P(x) \<Longrightarrow> Q) \<equiv> ((\<exists>x. P(x)) \<Longrightarrow> Q)\<close>
by rule iprover+
-lemma atomize_conjL[atomize_elim]: "(A \<Longrightarrow> B \<Longrightarrow> C) \<equiv> (A \<and> B \<Longrightarrow> C)"
+lemma atomize_conjL[atomize_elim]: \<open>(A \<Longrightarrow> B \<Longrightarrow> C) \<equiv> (A \<and> B \<Longrightarrow> C)\<close>
by rule iprover+
-lemma atomize_disjL[atomize_elim]: "((A \<Longrightarrow> C) \<Longrightarrow> (B \<Longrightarrow> C) \<Longrightarrow> C) \<equiv> ((A \<or> B \<Longrightarrow> C) \<Longrightarrow> C)"
+lemma atomize_disjL[atomize_elim]: \<open>((A \<Longrightarrow> C) \<Longrightarrow> (B \<Longrightarrow> C) \<Longrightarrow> C) \<equiv> ((A \<or> B \<Longrightarrow> C) \<Longrightarrow> C)\<close>
by rule iprover+
-lemma atomize_elimL[atomize_elim]: "(\<And>B. (A \<Longrightarrow> B) \<Longrightarrow> B) \<equiv> Trueprop(A)" ..
+lemma atomize_elimL[atomize_elim]: \<open>(\<And>B. (A \<Longrightarrow> B) \<Longrightarrow> B) \<equiv> Trueprop(A)\<close> ..
subsection \<open>Calculational rules\<close>
-lemma forw_subst: "a = b \<Longrightarrow> P(b) \<Longrightarrow> P(a)"
+lemma forw_subst: \<open>a = b \<Longrightarrow> P(b) \<Longrightarrow> P(a)\<close>
by (rule ssubst)
-lemma back_subst: "P(a) \<Longrightarrow> a = b \<Longrightarrow> P(b)"
+lemma back_subst: \<open>P(a) \<Longrightarrow> a = b \<Longrightarrow> P(b)\<close>
by (rule subst)
text \<open>
@@ -760,22 +760,22 @@
nonterminal letbinds and letbind
-definition Let :: "['a::{}, 'a => 'b] \<Rightarrow> ('b::{})"
- where "Let(s, f) \<equiv> f(s)"
+definition Let :: \<open>['a::{}, 'a => 'b] \<Rightarrow> ('b::{})\<close>
+ where \<open>Let(s, f) \<equiv> f(s)\<close>
syntax
- "_bind" :: "[pttrn, 'a] => letbind" (\<open>(2_ =/ _)\<close> 10)
- "" :: "letbind => letbinds" (\<open>_\<close>)
- "_binds" :: "[letbind, letbinds] => letbinds" (\<open>_;/ _\<close>)
- "_Let" :: "[letbinds, 'a] => 'a" (\<open>(let (_)/ in (_))\<close> 10)
+ "_bind" :: \<open>[pttrn, 'a] => letbind\<close> (\<open>(2_ =/ _)\<close> 10)
+ "" :: \<open>letbind => letbinds\<close> (\<open>_\<close>)
+ "_binds" :: \<open>[letbind, letbinds] => letbinds\<close> (\<open>_;/ _\<close>)
+ "_Let" :: \<open>[letbinds, 'a] => 'a\<close> (\<open>(let (_)/ in (_))\<close> 10)
translations
"_Let(_binds(b, bs), e)" == "_Let(b, _Let(bs, e))"
"let x = a in e" == "CONST Let(a, \<lambda>x. e)"
lemma LetI:
- assumes "\<And>x. x = t \<Longrightarrow> P(u(x))"
- shows "P(let x = t in u(x))"
+ assumes \<open>\<And>x. x = t \<Longrightarrow> P(u(x))\<close>
+ shows \<open>P(let x = t in u(x))\<close>
apply (unfold Let_def)
apply (rule refl [THEN assms])
done
@@ -784,113 +784,113 @@
subsection \<open>Intuitionistic simplification rules\<close>
lemma conj_simps:
- "P \<and> True \<longleftrightarrow> P"
- "True \<and> P \<longleftrightarrow> P"
- "P \<and> False \<longleftrightarrow> False"
- "False \<and> P \<longleftrightarrow> False"
- "P \<and> P \<longleftrightarrow> P"
- "P \<and> P \<and> Q \<longleftrightarrow> P \<and> Q"
- "P \<and> \<not> P \<longleftrightarrow> False"
- "\<not> P \<and> P \<longleftrightarrow> False"
- "(P \<and> Q) \<and> R \<longleftrightarrow> P \<and> (Q \<and> R)"
+ \<open>P \<and> True \<longleftrightarrow> P\<close>
+ \<open>True \<and> P \<longleftrightarrow> P\<close>
+ \<open>P \<and> False \<longleftrightarrow> False\<close>
+ \<open>False \<and> P \<longleftrightarrow> False\<close>
+ \<open>P \<and> P \<longleftrightarrow> P\<close>
+ \<open>P \<and> P \<and> Q \<longleftrightarrow> P \<and> Q\<close>
+ \<open>P \<and> \<not> P \<longleftrightarrow> False\<close>
+ \<open>\<not> P \<and> P \<longleftrightarrow> False\<close>
+ \<open>(P \<and> Q) \<and> R \<longleftrightarrow> P \<and> (Q \<and> R)\<close>
by iprover+
lemma disj_simps:
- "P \<or> True \<longleftrightarrow> True"
- "True \<or> P \<longleftrightarrow> True"
- "P \<or> False \<longleftrightarrow> P"
- "False \<or> P \<longleftrightarrow> P"
- "P \<or> P \<longleftrightarrow> P"
- "P \<or> P \<or> Q \<longleftrightarrow> P \<or> Q"
- "(P \<or> Q) \<or> R \<longleftrightarrow> P \<or> (Q \<or> R)"
+ \<open>P \<or> True \<longleftrightarrow> True\<close>
+ \<open>True \<or> P \<longleftrightarrow> True\<close>
+ \<open>P \<or> False \<longleftrightarrow> P\<close>
+ \<open>False \<or> P \<longleftrightarrow> P\<close>
+ \<open>P \<or> P \<longleftrightarrow> P\<close>
+ \<open>P \<or> P \<or> Q \<longleftrightarrow> P \<or> Q\<close>
+ \<open>(P \<or> Q) \<or> R \<longleftrightarrow> P \<or> (Q \<or> R)\<close>
by iprover+
lemma not_simps:
- "\<not> (P \<or> Q) \<longleftrightarrow> \<not> P \<and> \<not> Q"
- "\<not> False \<longleftrightarrow> True"
- "\<not> True \<longleftrightarrow> False"
+ \<open>\<not> (P \<or> Q) \<longleftrightarrow> \<not> P \<and> \<not> Q\<close>
+ \<open>\<not> False \<longleftrightarrow> True\<close>
+ \<open>\<not> True \<longleftrightarrow> False\<close>
by iprover+
lemma imp_simps:
- "(P \<longrightarrow> False) \<longleftrightarrow> \<not> P"
- "(P \<longrightarrow> True) \<longleftrightarrow> True"
- "(False \<longrightarrow> P) \<longleftrightarrow> True"
- "(True \<longrightarrow> P) \<longleftrightarrow> P"
- "(P \<longrightarrow> P) \<longleftrightarrow> True"
- "(P \<longrightarrow> \<not> P) \<longleftrightarrow> \<not> P"
+ \<open>(P \<longrightarrow> False) \<longleftrightarrow> \<not> P\<close>
+ \<open>(P \<longrightarrow> True) \<longleftrightarrow> True\<close>
+ \<open>(False \<longrightarrow> P) \<longleftrightarrow> True\<close>
+ \<open>(True \<longrightarrow> P) \<longleftrightarrow> P\<close>
+ \<open>(P \<longrightarrow> P) \<longleftrightarrow> True\<close>
+ \<open>(P \<longrightarrow> \<not> P) \<longleftrightarrow> \<not> P\<close>
by iprover+
lemma iff_simps:
- "(True \<longleftrightarrow> P) \<longleftrightarrow> P"
- "(P \<longleftrightarrow> True) \<longleftrightarrow> P"
- "(P \<longleftrightarrow> P) \<longleftrightarrow> True"
- "(False \<longleftrightarrow> P) \<longleftrightarrow> \<not> P"
- "(P \<longleftrightarrow> False) \<longleftrightarrow> \<not> P"
+ \<open>(True \<longleftrightarrow> P) \<longleftrightarrow> P\<close>
+ \<open>(P \<longleftrightarrow> True) \<longleftrightarrow> P\<close>
+ \<open>(P \<longleftrightarrow> P) \<longleftrightarrow> True\<close>
+ \<open>(False \<longleftrightarrow> P) \<longleftrightarrow> \<not> P\<close>
+ \<open>(P \<longleftrightarrow> False) \<longleftrightarrow> \<not> P\<close>
by iprover+
text \<open>The \<open>x = t\<close> versions are needed for the simplification
procedures.\<close>
lemma quant_simps:
- "\<And>P. (\<forall>x. P) \<longleftrightarrow> P"
- "(\<forall>x. x = t \<longrightarrow> P(x)) \<longleftrightarrow> P(t)"
- "(\<forall>x. t = x \<longrightarrow> P(x)) \<longleftrightarrow> P(t)"
- "\<And>P. (\<exists>x. P) \<longleftrightarrow> P"
- "\<exists>x. x = t"
- "\<exists>x. t = x"
- "(\<exists>x. x = t \<and> P(x)) \<longleftrightarrow> P(t)"
- "(\<exists>x. t = x \<and> P(x)) \<longleftrightarrow> P(t)"
+ \<open>\<And>P. (\<forall>x. P) \<longleftrightarrow> P\<close>
+ \<open>(\<forall>x. x = t \<longrightarrow> P(x)) \<longleftrightarrow> P(t)\<close>
+ \<open>(\<forall>x. t = x \<longrightarrow> P(x)) \<longleftrightarrow> P(t)\<close>
+ \<open>\<And>P. (\<exists>x. P) \<longleftrightarrow> P\<close>
+ \<open>\<exists>x. x = t\<close>
+ \<open>\<exists>x. t = x\<close>
+ \<open>(\<exists>x. x = t \<and> P(x)) \<longleftrightarrow> P(t)\<close>
+ \<open>(\<exists>x. t = x \<and> P(x)) \<longleftrightarrow> P(t)\<close>
by iprover+
text \<open>These are NOT supplied by default!\<close>
lemma distrib_simps:
- "P \<and> (Q \<or> R) \<longleftrightarrow> P \<and> Q \<or> P \<and> R"
- "(Q \<or> R) \<and> P \<longleftrightarrow> Q \<and> P \<or> R \<and> P"
- "(P \<or> Q \<longrightarrow> R) \<longleftrightarrow> (P \<longrightarrow> R) \<and> (Q \<longrightarrow> R)"
+ \<open>P \<and> (Q \<or> R) \<longleftrightarrow> P \<and> Q \<or> P \<and> R\<close>
+ \<open>(Q \<or> R) \<and> P \<longleftrightarrow> Q \<and> P \<or> R \<and> P\<close>
+ \<open>(P \<or> Q \<longrightarrow> R) \<longleftrightarrow> (P \<longrightarrow> R) \<and> (Q \<longrightarrow> R)\<close>
by iprover+
subsubsection \<open>Conversion into rewrite rules\<close>
-lemma P_iff_F: "\<not> P \<Longrightarrow> (P \<longleftrightarrow> False)"
+lemma P_iff_F: \<open>\<not> P \<Longrightarrow> (P \<longleftrightarrow> False)\<close>
by iprover
-lemma iff_reflection_F: "\<not> P \<Longrightarrow> (P \<equiv> False)"
+lemma iff_reflection_F: \<open>\<not> P \<Longrightarrow> (P \<equiv> False)\<close>
by (rule P_iff_F [THEN iff_reflection])
-lemma P_iff_T: "P \<Longrightarrow> (P \<longleftrightarrow> True)"
+lemma P_iff_T: \<open>P \<Longrightarrow> (P \<longleftrightarrow> True)\<close>
by iprover
-lemma iff_reflection_T: "P \<Longrightarrow> (P \<equiv> True)"
+lemma iff_reflection_T: \<open>P \<Longrightarrow> (P \<equiv> True)\<close>
by (rule P_iff_T [THEN iff_reflection])
subsubsection \<open>More rewrite rules\<close>
-lemma conj_commute: "P \<and> Q \<longleftrightarrow> Q \<and> P" by iprover
-lemma conj_left_commute: "P \<and> (Q \<and> R) \<longleftrightarrow> Q \<and> (P \<and> R)" by iprover
+lemma conj_commute: \<open>P \<and> Q \<longleftrightarrow> Q \<and> P\<close> by iprover
+lemma conj_left_commute: \<open>P \<and> (Q \<and> R) \<longleftrightarrow> Q \<and> (P \<and> R)\<close> by iprover
lemmas conj_comms = conj_commute conj_left_commute
-lemma disj_commute: "P \<or> Q \<longleftrightarrow> Q \<or> P" by iprover
-lemma disj_left_commute: "P \<or> (Q \<or> R) \<longleftrightarrow> Q \<or> (P \<or> R)" by iprover
+lemma disj_commute: \<open>P \<or> Q \<longleftrightarrow> Q \<or> P\<close> by iprover
+lemma disj_left_commute: \<open>P \<or> (Q \<or> R) \<longleftrightarrow> Q \<or> (P \<or> R)\<close> by iprover
lemmas disj_comms = disj_commute disj_left_commute
-lemma conj_disj_distribL: "P \<and> (Q \<or> R) \<longleftrightarrow> (P \<and> Q \<or> P \<and> R)" by iprover
-lemma conj_disj_distribR: "(P \<or> Q) \<and> R \<longleftrightarrow> (P \<and> R \<or> Q \<and> R)" by iprover
+lemma conj_disj_distribL: \<open>P \<and> (Q \<or> R) \<longleftrightarrow> (P \<and> Q \<or> P \<and> R)\<close> by iprover
+lemma conj_disj_distribR: \<open>(P \<or> Q) \<and> R \<longleftrightarrow> (P \<and> R \<or> Q \<and> R)\<close> by iprover
-lemma disj_conj_distribL: "P \<or> (Q \<and> R) \<longleftrightarrow> (P \<or> Q) \<and> (P \<or> R)" by iprover
-lemma disj_conj_distribR: "(P \<and> Q) \<or> R \<longleftrightarrow> (P \<or> R) \<and> (Q \<or> R)" by iprover
+lemma disj_conj_distribL: \<open>P \<or> (Q \<and> R) \<longleftrightarrow> (P \<or> Q) \<and> (P \<or> R)\<close> by iprover
+lemma disj_conj_distribR: \<open>(P \<and> Q) \<or> R \<longleftrightarrow> (P \<or> R) \<and> (Q \<or> R)\<close> by iprover
-lemma imp_conj_distrib: "(P \<longrightarrow> (Q \<and> R)) \<longleftrightarrow> (P \<longrightarrow> Q) \<and> (P \<longrightarrow> R)" by iprover
-lemma imp_conj: "((P \<and> Q) \<longrightarrow> R) \<longleftrightarrow> (P \<longrightarrow> (Q \<longrightarrow> R))" by iprover
-lemma imp_disj: "(P \<or> Q \<longrightarrow> R) \<longleftrightarrow> (P \<longrightarrow> R) \<and> (Q \<longrightarrow> R)" by iprover
+lemma imp_conj_distrib: \<open>(P \<longrightarrow> (Q \<and> R)) \<longleftrightarrow> (P \<longrightarrow> Q) \<and> (P \<longrightarrow> R)\<close> by iprover
+lemma imp_conj: \<open>((P \<and> Q) \<longrightarrow> R) \<longleftrightarrow> (P \<longrightarrow> (Q \<longrightarrow> R))\<close> by iprover
+lemma imp_disj: \<open>(P \<or> Q \<longrightarrow> R) \<longleftrightarrow> (P \<longrightarrow> R) \<and> (Q \<longrightarrow> R)\<close> by iprover
-lemma de_Morgan_disj: "(\<not> (P \<or> Q)) \<longleftrightarrow> (\<not> P \<and> \<not> Q)" by iprover
+lemma de_Morgan_disj: \<open>(\<not> (P \<or> Q)) \<longleftrightarrow> (\<not> P \<and> \<not> Q)\<close> by iprover
-lemma not_ex: "(\<not> (\<exists>x. P(x))) \<longleftrightarrow> (\<forall>x. \<not> P(x))" by iprover
-lemma imp_ex: "((\<exists>x. P(x)) \<longrightarrow> Q) \<longleftrightarrow> (\<forall>x. P(x) \<longrightarrow> Q)" by iprover
+lemma not_ex: \<open>(\<not> (\<exists>x. P(x))) \<longleftrightarrow> (\<forall>x. \<not> P(x))\<close> by iprover
+lemma imp_ex: \<open>((\<exists>x. P(x)) \<longrightarrow> Q) \<longleftrightarrow> (\<forall>x. P(x) \<longrightarrow> Q)\<close> by iprover
-lemma ex_disj_distrib: "(\<exists>x. P(x) \<or> Q(x)) \<longleftrightarrow> ((\<exists>x. P(x)) \<or> (\<exists>x. Q(x)))"
+lemma ex_disj_distrib: \<open>(\<exists>x. P(x) \<or> Q(x)) \<longleftrightarrow> ((\<exists>x. P(x)) \<or> (\<exists>x. Q(x)))\<close>
by iprover
-lemma all_conj_distrib: "(\<forall>x. P(x) \<and> Q(x)) \<longleftrightarrow> ((\<forall>x. P(x)) \<and> (\<forall>x. Q(x)))"
+lemma all_conj_distrib: \<open>(\<forall>x. P(x) \<and> Q(x)) \<longleftrightarrow> ((\<forall>x. P(x)) \<and> (\<forall>x. Q(x)))\<close>
by iprover
end
--- a/src/FOL/ex/Classical.thy Thu Jan 03 21:48:05 2019 +0100
+++ b/src/FOL/ex/Classical.thy Thu Jan 03 22:19:19 2019 +0100
@@ -9,16 +9,16 @@
imports FOL
begin
-lemma "(P \<longrightarrow> Q \<or> R) \<longrightarrow> (P \<longrightarrow> Q) \<or> (P \<longrightarrow> R)"
+lemma \<open>(P \<longrightarrow> Q \<or> R) \<longrightarrow> (P \<longrightarrow> Q) \<or> (P \<longrightarrow> R)\<close>
by blast
subsubsection \<open>If and only if\<close>
-lemma "(P \<longleftrightarrow> Q) \<longleftrightarrow> (Q \<longleftrightarrow> P)"
+lemma \<open>(P \<longleftrightarrow> Q) \<longleftrightarrow> (Q \<longleftrightarrow> P)\<close>
by blast
-lemma "\<not> (P \<longleftrightarrow> \<not> P)"
+lemma \<open>\<not> (P \<longleftrightarrow> \<not> P)\<close>
by blast
@@ -37,91 +37,91 @@
\<close>
text\<open>1\<close>
-lemma "(P \<longrightarrow> Q) \<longleftrightarrow> (\<not> Q \<longrightarrow> \<not> P)"
+lemma \<open>(P \<longrightarrow> Q) \<longleftrightarrow> (\<not> Q \<longrightarrow> \<not> P)\<close>
by blast
text\<open>2\<close>
-lemma "\<not> \<not> P \<longleftrightarrow> P"
+lemma \<open>\<not> \<not> P \<longleftrightarrow> P\<close>
by blast
text\<open>3\<close>
-lemma "\<not> (P \<longrightarrow> Q) \<longrightarrow> (Q \<longrightarrow> P)"
+lemma \<open>\<not> (P \<longrightarrow> Q) \<longrightarrow> (Q \<longrightarrow> P)\<close>
by blast
text\<open>4\<close>
-lemma "(\<not> P \<longrightarrow> Q) \<longleftrightarrow> (\<not> Q \<longrightarrow> P)"
+lemma \<open>(\<not> P \<longrightarrow> Q) \<longleftrightarrow> (\<not> Q \<longrightarrow> P)\<close>
by blast
text\<open>5\<close>
-lemma "((P \<or> Q) \<longrightarrow> (P \<or> R)) \<longrightarrow> (P \<or> (Q \<longrightarrow> R))"
+lemma \<open>((P \<or> Q) \<longrightarrow> (P \<or> R)) \<longrightarrow> (P \<or> (Q \<longrightarrow> R))\<close>
by blast
text\<open>6\<close>
-lemma "P \<or> \<not> P"
+lemma \<open>P \<or> \<not> P\<close>
by blast
text\<open>7\<close>
-lemma "P \<or> \<not> \<not> \<not> P"
+lemma \<open>P \<or> \<not> \<not> \<not> P\<close>
by blast
text\<open>8. Peirce's law\<close>
-lemma "((P \<longrightarrow> Q) \<longrightarrow> P) \<longrightarrow> P"
+lemma \<open>((P \<longrightarrow> Q) \<longrightarrow> P) \<longrightarrow> P\<close>
by blast
text\<open>9\<close>
-lemma "((P \<or> Q) \<and> (\<not> P \<or> Q) \<and> (P \<or> \<not> Q)) \<longrightarrow> \<not> (\<not> P \<or> \<not> Q)"
+lemma \<open>((P \<or> Q) \<and> (\<not> P \<or> Q) \<and> (P \<or> \<not> Q)) \<longrightarrow> \<not> (\<not> P \<or> \<not> Q)\<close>
by blast
text\<open>10\<close>
-lemma "(Q \<longrightarrow> R) \<and> (R \<longrightarrow> P \<and> Q) \<and> (P \<longrightarrow> Q \<or> R) \<longrightarrow> (P \<longleftrightarrow> Q)"
+lemma \<open>(Q \<longrightarrow> R) \<and> (R \<longrightarrow> P \<and> Q) \<and> (P \<longrightarrow> Q \<or> R) \<longrightarrow> (P \<longleftrightarrow> Q)\<close>
by blast
text\<open>11. Proved in each direction (incorrectly, says Pelletier!!)\<close>
-lemma "P \<longleftrightarrow> P"
+lemma \<open>P \<longleftrightarrow> P\<close>
by blast
text\<open>12. "Dijkstra's law"\<close>
-lemma "((P \<longleftrightarrow> Q) \<longleftrightarrow> R) \<longleftrightarrow> (P \<longleftrightarrow> (Q \<longleftrightarrow> R))"
+lemma \<open>((P \<longleftrightarrow> Q) \<longleftrightarrow> R) \<longleftrightarrow> (P \<longleftrightarrow> (Q \<longleftrightarrow> R))\<close>
by blast
text\<open>13. Distributive law\<close>
-lemma "P \<or> (Q \<and> R) \<longleftrightarrow> (P \<or> Q) \<and> (P \<or> R)"
+lemma \<open>P \<or> (Q \<and> R) \<longleftrightarrow> (P \<or> Q) \<and> (P \<or> R)\<close>
by blast
text\<open>14\<close>
-lemma "(P \<longleftrightarrow> Q) \<longleftrightarrow> ((Q \<or> \<not> P) \<and> (\<not> Q \<or> P))"
+lemma \<open>(P \<longleftrightarrow> Q) \<longleftrightarrow> ((Q \<or> \<not> P) \<and> (\<not> Q \<or> P))\<close>
by blast
text\<open>15\<close>
-lemma "(P \<longrightarrow> Q) \<longleftrightarrow> (\<not> P \<or> Q)"
+lemma \<open>(P \<longrightarrow> Q) \<longleftrightarrow> (\<not> P \<or> Q)\<close>
by blast
text\<open>16\<close>
-lemma "(P \<longrightarrow> Q) \<or> (Q \<longrightarrow> P)"
+lemma \<open>(P \<longrightarrow> Q) \<or> (Q \<longrightarrow> P)\<close>
by blast
text\<open>17\<close>
-lemma "((P \<and> (Q \<longrightarrow> R)) \<longrightarrow> S) \<longleftrightarrow> ((\<not> P \<or> Q \<or> S) \<and> (\<not> P \<or> \<not> R \<or> S))"
+lemma \<open>((P \<and> (Q \<longrightarrow> R)) \<longrightarrow> S) \<longleftrightarrow> ((\<not> P \<or> Q \<or> S) \<and> (\<not> P \<or> \<not> R \<or> S))\<close>
by blast
subsection \<open>Classical Logic: examples with quantifiers\<close>
-lemma "(\<forall>x. P(x) \<and> Q(x)) \<longleftrightarrow> (\<forall>x. P(x)) \<and> (\<forall>x. Q(x))"
+lemma \<open>(\<forall>x. P(x) \<and> Q(x)) \<longleftrightarrow> (\<forall>x. P(x)) \<and> (\<forall>x. Q(x))\<close>
by blast
-lemma "(\<exists>x. P \<longrightarrow> Q(x)) \<longleftrightarrow> (P \<longrightarrow> (\<exists>x. Q(x)))"
+lemma \<open>(\<exists>x. P \<longrightarrow> Q(x)) \<longleftrightarrow> (P \<longrightarrow> (\<exists>x. Q(x)))\<close>
by blast
-lemma "(\<exists>x. P(x) \<longrightarrow> Q) \<longleftrightarrow> (\<forall>x. P(x)) \<longrightarrow> Q"
+lemma \<open>(\<exists>x. P(x) \<longrightarrow> Q) \<longleftrightarrow> (\<forall>x. P(x)) \<longrightarrow> Q\<close>
by blast
-lemma "(\<forall>x. P(x)) \<or> Q \<longleftrightarrow> (\<forall>x. P(x) \<or> Q)"
+lemma \<open>(\<forall>x. P(x)) \<or> Q \<longleftrightarrow> (\<forall>x. P(x) \<or> Q)\<close>
by blast
text\<open>Discussed in Avron, Gentzen-Type Systems, Resolution and Tableaux,
JAR 10 (265-281), 1993. Proof is trivial!\<close>
-lemma "\<not> ((\<exists>x. \<not> P(x)) \<and> ((\<exists>x. P(x)) \<or> (\<exists>x. P(x) \<and> Q(x))) \<and> \<not> (\<exists>x. P(x)))"
+lemma \<open>\<not> ((\<exists>x. \<not> P(x)) \<and> ((\<exists>x. P(x)) \<or> (\<exists>x. P(x) \<and> Q(x))) \<and> \<not> (\<exists>x. P(x)))\<close>
by blast
@@ -129,195 +129,195 @@
text\<open>Theorem B of Peter Andrews, Theorem Proving via General Matings,
JACM 28 (1981).\<close>
-lemma "(\<exists>x. \<forall>y. P(x) \<longleftrightarrow> P(y)) \<longrightarrow> ((\<exists>x. P(x)) \<longleftrightarrow> (\<forall>y. P(y)))"
+lemma \<open>(\<exists>x. \<forall>y. P(x) \<longleftrightarrow> P(y)) \<longrightarrow> ((\<exists>x. P(x)) \<longleftrightarrow> (\<forall>y. P(y)))\<close>
by blast
text\<open>Needs multiple instantiation of ALL.\<close>
-lemma "(\<forall>x. P(x) \<longrightarrow> P(f(x))) \<and> P(d) \<longrightarrow> P(f(f(f(d))))"
+lemma \<open>(\<forall>x. P(x) \<longrightarrow> P(f(x))) \<and> P(d) \<longrightarrow> P(f(f(f(d))))\<close>
by blast
text\<open>Needs double instantiation of the quantifier\<close>
-lemma "\<exists>x. P(x) \<longrightarrow> P(a) \<and> P(b)"
+lemma \<open>\<exists>x. P(x) \<longrightarrow> P(a) \<and> P(b)\<close>
by blast
-lemma "\<exists>z. P(z) \<longrightarrow> (\<forall>x. P(x))"
+lemma \<open>\<exists>z. P(z) \<longrightarrow> (\<forall>x. P(x))\<close>
by blast
-lemma "\<exists>x. (\<exists>y. P(y)) \<longrightarrow> P(x)"
+lemma \<open>\<exists>x. (\<exists>y. P(y)) \<longrightarrow> P(x)\<close>
by blast
text\<open>V. Lifschitz, What Is the Inverse Method?, JAR 5 (1989), 1--23. NOT PROVED.\<close>
lemma
- "\<exists>x x'. \<forall>y. \<exists>z z'.
+ \<open>\<exists>x x'. \<forall>y. \<exists>z z'.
(\<not> P(y,y) \<or> P(x,x) \<or> \<not> S(z,x)) \<and>
(S(x,y) \<or> \<not> S(y,z) \<or> Q(z',z')) \<and>
- (Q(x',y) \<or> \<not> Q(y,z') \<or> S(x',x'))"
+ (Q(x',y) \<or> \<not> Q(y,z') \<or> S(x',x'))\<close>
oops
subsection \<open>Hard examples with quantifiers\<close>
text\<open>18\<close>
-lemma "\<exists>y. \<forall>x. P(y) \<longrightarrow> P(x)"
+lemma \<open>\<exists>y. \<forall>x. P(y) \<longrightarrow> P(x)\<close>
by blast
text\<open>19\<close>
-lemma "\<exists>x. \<forall>y z. (P(y) \<longrightarrow> Q(z)) \<longrightarrow> (P(x) \<longrightarrow> Q(x))"
+lemma \<open>\<exists>x. \<forall>y z. (P(y) \<longrightarrow> Q(z)) \<longrightarrow> (P(x) \<longrightarrow> Q(x))\<close>
by blast
text\<open>20\<close>
-lemma "(\<forall>x y. \<exists>z. \<forall>w. (P(x) \<and> Q(y) \<longrightarrow> R(z) \<and> S(w)))
- \<longrightarrow> (\<exists>x y. P(x) \<and> Q(y)) \<longrightarrow> (\<exists>z. R(z))"
+lemma \<open>(\<forall>x y. \<exists>z. \<forall>w. (P(x) \<and> Q(y) \<longrightarrow> R(z) \<and> S(w)))
+ \<longrightarrow> (\<exists>x y. P(x) \<and> Q(y)) \<longrightarrow> (\<exists>z. R(z))\<close>
by blast
text\<open>21\<close>
-lemma "(\<exists>x. P \<longrightarrow> Q(x)) \<and> (\<exists>x. Q(x) \<longrightarrow> P) \<longrightarrow> (\<exists>x. P \<longleftrightarrow> Q(x))"
+lemma \<open>(\<exists>x. P \<longrightarrow> Q(x)) \<and> (\<exists>x. Q(x) \<longrightarrow> P) \<longrightarrow> (\<exists>x. P \<longleftrightarrow> Q(x))\<close>
by blast
text\<open>22\<close>
-lemma "(\<forall>x. P \<longleftrightarrow> Q(x)) \<longrightarrow> (P \<longleftrightarrow> (\<forall>x. Q(x)))"
+lemma \<open>(\<forall>x. P \<longleftrightarrow> Q(x)) \<longrightarrow> (P \<longleftrightarrow> (\<forall>x. Q(x)))\<close>
by blast
text\<open>23\<close>
-lemma "(\<forall>x. P \<or> Q(x)) \<longleftrightarrow> (P \<or> (\<forall>x. Q(x)))"
+lemma \<open>(\<forall>x. P \<or> Q(x)) \<longleftrightarrow> (P \<or> (\<forall>x. Q(x)))\<close>
by blast
text\<open>24\<close>
lemma
- "\<not> (\<exists>x. S(x) \<and> Q(x)) \<and> (\<forall>x. P(x) \<longrightarrow> Q(x) \<or> R(x)) \<and>
+ \<open>\<not> (\<exists>x. S(x) \<and> Q(x)) \<and> (\<forall>x. P(x) \<longrightarrow> Q(x) \<or> R(x)) \<and>
(\<not> (\<exists>x. P(x)) \<longrightarrow> (\<exists>x. Q(x))) \<and> (\<forall>x. Q(x) \<or> R(x) \<longrightarrow> S(x))
- \<longrightarrow> (\<exists>x. P(x) \<and> R(x))"
+ \<longrightarrow> (\<exists>x. P(x) \<and> R(x))\<close>
by blast
text\<open>25\<close>
lemma
- "(\<exists>x. P(x)) \<and>
+ \<open>(\<exists>x. P(x)) \<and>
(\<forall>x. L(x) \<longrightarrow> \<not> (M(x) \<and> R(x))) \<and>
(\<forall>x. P(x) \<longrightarrow> (M(x) \<and> L(x))) \<and>
((\<forall>x. P(x) \<longrightarrow> Q(x)) \<or> (\<exists>x. P(x) \<and> R(x)))
- \<longrightarrow> (\<exists>x. Q(x) \<and> P(x))"
+ \<longrightarrow> (\<exists>x. Q(x) \<and> P(x))\<close>
by blast
text\<open>26\<close>
lemma
- "((\<exists>x. p(x)) \<longleftrightarrow> (\<exists>x. q(x))) \<and>
+ \<open>((\<exists>x. p(x)) \<longleftrightarrow> (\<exists>x. q(x))) \<and>
(\<forall>x. \<forall>y. p(x) \<and> q(y) \<longrightarrow> (r(x) \<longleftrightarrow> s(y)))
- \<longrightarrow> ((\<forall>x. p(x) \<longrightarrow> r(x)) \<longleftrightarrow> (\<forall>x. q(x) \<longrightarrow> s(x)))"
+ \<longrightarrow> ((\<forall>x. p(x) \<longrightarrow> r(x)) \<longleftrightarrow> (\<forall>x. q(x) \<longrightarrow> s(x)))\<close>
by blast
text\<open>27\<close>
lemma
- "(\<exists>x. P(x) \<and> \<not> Q(x)) \<and>
+ \<open>(\<exists>x. P(x) \<and> \<not> Q(x)) \<and>
(\<forall>x. P(x) \<longrightarrow> R(x)) \<and>
(\<forall>x. M(x) \<and> L(x) \<longrightarrow> P(x)) \<and>
((\<exists>x. R(x) \<and> \<not> Q(x)) \<longrightarrow> (\<forall>x. L(x) \<longrightarrow> \<not> R(x)))
- \<longrightarrow> (\<forall>x. M(x) \<longrightarrow> \<not> L(x))"
+ \<longrightarrow> (\<forall>x. M(x) \<longrightarrow> \<not> L(x))\<close>
by blast
text\<open>28. AMENDED\<close>
lemma
- "(\<forall>x. P(x) \<longrightarrow> (\<forall>x. Q(x))) \<and>
+ \<open>(\<forall>x. P(x) \<longrightarrow> (\<forall>x. Q(x))) \<and>
((\<forall>x. Q(x) \<or> R(x)) \<longrightarrow> (\<exists>x. Q(x) \<and> S(x))) \<and>
((\<exists>x. S(x)) \<longrightarrow> (\<forall>x. L(x) \<longrightarrow> M(x)))
- \<longrightarrow> (\<forall>x. P(x) \<and> L(x) \<longrightarrow> M(x))"
+ \<longrightarrow> (\<forall>x. P(x) \<and> L(x) \<longrightarrow> M(x))\<close>
by blast
text\<open>29. Essentially the same as Principia Mathematica *11.71\<close>
lemma
- "(\<exists>x. P(x)) \<and> (\<exists>y. Q(y))
+ \<open>(\<exists>x. P(x)) \<and> (\<exists>y. Q(y))
\<longrightarrow> ((\<forall>x. P(x) \<longrightarrow> R(x)) \<and> (\<forall>y. Q(y) \<longrightarrow> S(y)) \<longleftrightarrow>
- (\<forall>x y. P(x) \<and> Q(y) \<longrightarrow> R(x) \<and> S(y)))"
+ (\<forall>x y. P(x) \<and> Q(y) \<longrightarrow> R(x) \<and> S(y)))\<close>
by blast
text\<open>30\<close>
lemma
- "(\<forall>x. P(x) \<or> Q(x) \<longrightarrow> \<not> R(x)) \<and>
+ \<open>(\<forall>x. P(x) \<or> Q(x) \<longrightarrow> \<not> R(x)) \<and>
(\<forall>x. (Q(x) \<longrightarrow> \<not> S(x)) \<longrightarrow> P(x) \<and> R(x))
- \<longrightarrow> (\<forall>x. S(x))"
+ \<longrightarrow> (\<forall>x. S(x))\<close>
by blast
text\<open>31\<close>
lemma
- "\<not> (\<exists>x. P(x) \<and> (Q(x) \<or> R(x))) \<and>
+ \<open>\<not> (\<exists>x. P(x) \<and> (Q(x) \<or> R(x))) \<and>
(\<exists>x. L(x) \<and> P(x)) \<and>
(\<forall>x. \<not> R(x) \<longrightarrow> M(x))
- \<longrightarrow> (\<exists>x. L(x) \<and> M(x))"
+ \<longrightarrow> (\<exists>x. L(x) \<and> M(x))\<close>
by blast
text\<open>32\<close>
lemma
- "(\<forall>x. P(x) \<and> (Q(x) \<or> R(x)) \<longrightarrow> S(x)) \<and>
+ \<open>(\<forall>x. P(x) \<and> (Q(x) \<or> R(x)) \<longrightarrow> S(x)) \<and>
(\<forall>x. S(x) \<and> R(x) \<longrightarrow> L(x)) \<and>
(\<forall>x. M(x) \<longrightarrow> R(x))
- \<longrightarrow> (\<forall>x. P(x) \<and> M(x) \<longrightarrow> L(x))"
+ \<longrightarrow> (\<forall>x. P(x) \<and> M(x) \<longrightarrow> L(x))\<close>
by blast
text\<open>33\<close>
lemma
- "(\<forall>x. P(a) \<and> (P(x) \<longrightarrow> P(b)) \<longrightarrow> P(c)) \<longleftrightarrow>
- (\<forall>x. (\<not> P(a) \<or> P(x) \<or> P(c)) \<and> (\<not> P(a) \<or> \<not> P(b) \<or> P(c)))"
+ \<open>(\<forall>x. P(a) \<and> (P(x) \<longrightarrow> P(b)) \<longrightarrow> P(c)) \<longleftrightarrow>
+ (\<forall>x. (\<not> P(a) \<or> P(x) \<or> P(c)) \<and> (\<not> P(a) \<or> \<not> P(b) \<or> P(c)))\<close>
by blast
text\<open>34. AMENDED (TWICE!!). Andrews's challenge.\<close>
lemma
- "((\<exists>x. \<forall>y. p(x) \<longleftrightarrow> p(y)) \<longleftrightarrow> ((\<exists>x. q(x)) \<longleftrightarrow> (\<forall>y. p(y)))) \<longleftrightarrow>
- ((\<exists>x. \<forall>y. q(x) \<longleftrightarrow> q(y)) \<longleftrightarrow> ((\<exists>x. p(x)) \<longleftrightarrow> (\<forall>y. q(y))))"
+ \<open>((\<exists>x. \<forall>y. p(x) \<longleftrightarrow> p(y)) \<longleftrightarrow> ((\<exists>x. q(x)) \<longleftrightarrow> (\<forall>y. p(y)))) \<longleftrightarrow>
+ ((\<exists>x. \<forall>y. q(x) \<longleftrightarrow> q(y)) \<longleftrightarrow> ((\<exists>x. p(x)) \<longleftrightarrow> (\<forall>y. q(y))))\<close>
by blast
text\<open>35\<close>
-lemma "\<exists>x y. P(x,y) \<longrightarrow> (\<forall>u v. P(u,v))"
+lemma \<open>\<exists>x y. P(x,y) \<longrightarrow> (\<forall>u v. P(u,v))\<close>
by blast
text\<open>36\<close>
lemma
- "(\<forall>x. \<exists>y. J(x,y)) \<and>
+ \<open>(\<forall>x. \<exists>y. J(x,y)) \<and>
(\<forall>x. \<exists>y. G(x,y)) \<and>
(\<forall>x y. J(x,y) \<or> G(x,y) \<longrightarrow> (\<forall>z. J(y,z) \<or> G(y,z) \<longrightarrow> H(x,z)))
- \<longrightarrow> (\<forall>x. \<exists>y. H(x,y))"
+ \<longrightarrow> (\<forall>x. \<exists>y. H(x,y))\<close>
by blast
text\<open>37\<close>
lemma
- "(\<forall>z. \<exists>w. \<forall>x. \<exists>y.
+ \<open>(\<forall>z. \<exists>w. \<forall>x. \<exists>y.
(P(x,z) \<longrightarrow> P(y,w)) \<and> P(y,z) \<and> (P(y,w) \<longrightarrow> (\<exists>u. Q(u,w)))) \<and>
(\<forall>x z. \<not> P(x,z) \<longrightarrow> (\<exists>y. Q(y,z))) \<and>
((\<exists>x y. Q(x,y)) \<longrightarrow> (\<forall>x. R(x,x)))
- \<longrightarrow> (\<forall>x. \<exists>y. R(x,y))"
+ \<longrightarrow> (\<forall>x. \<exists>y. R(x,y))\<close>
by blast
text\<open>38\<close>
lemma
- "(\<forall>x. p(a) \<and> (p(x) \<longrightarrow> (\<exists>y. p(y) \<and> r(x,y))) \<longrightarrow>
+ \<open>(\<forall>x. p(a) \<and> (p(x) \<longrightarrow> (\<exists>y. p(y) \<and> r(x,y))) \<longrightarrow>
(\<exists>z. \<exists>w. p(z) \<and> r(x,w) \<and> r(w,z))) \<longleftrightarrow>
(\<forall>x. (\<not> p(a) \<or> p(x) \<or> (\<exists>z. \<exists>w. p(z) \<and> r(x,w) \<and> r(w,z))) \<and>
(\<not> p(a) \<or> \<not> (\<exists>y. p(y) \<and> r(x,y)) \<or>
- (\<exists>z. \<exists>w. p(z) \<and> r(x,w) \<and> r(w,z))))"
+ (\<exists>z. \<exists>w. p(z) \<and> r(x,w) \<and> r(w,z))))\<close>
by blast
text\<open>39\<close>
-lemma "\<not> (\<exists>x. \<forall>y. F(y,x) \<longleftrightarrow> \<not> F(y,y))"
+lemma \<open>\<not> (\<exists>x. \<forall>y. F(y,x) \<longleftrightarrow> \<not> F(y,y))\<close>
by blast
text\<open>40. AMENDED\<close>
lemma
- "(\<exists>y. \<forall>x. F(x,y) \<longleftrightarrow> F(x,x)) \<longrightarrow>
- \<not> (\<forall>x. \<exists>y. \<forall>z. F(z,y) \<longleftrightarrow> \<not> F(z,x))"
+ \<open>(\<exists>y. \<forall>x. F(x,y) \<longleftrightarrow> F(x,x)) \<longrightarrow>
+ \<not> (\<forall>x. \<exists>y. \<forall>z. F(z,y) \<longleftrightarrow> \<not> F(z,x))\<close>
by blast
text\<open>41\<close>
lemma
- "(\<forall>z. \<exists>y. \<forall>x. f(x,y) \<longleftrightarrow> f(x,z) \<and> \<not> f(x,x))
- \<longrightarrow> \<not> (\<exists>z. \<forall>x. f(x,z))"
+ \<open>(\<forall>z. \<exists>y. \<forall>x. f(x,y) \<longleftrightarrow> f(x,z) \<and> \<not> f(x,x))
+ \<longrightarrow> \<not> (\<exists>z. \<forall>x. f(x,z))\<close>
by blast
text\<open>42\<close>
-lemma "\<not> (\<exists>y. \<forall>x. p(x,y) \<longleftrightarrow> \<not> (\<exists>z. p(x,z) \<and> p(z,x)))"
+lemma \<open>\<not> (\<exists>y. \<forall>x. p(x,y) \<longleftrightarrow> \<not> (\<exists>z. p(x,z) \<and> p(z,x)))\<close>
by blast
text\<open>43\<close>
lemma
- "(\<forall>x. \<forall>y. q(x,y) \<longleftrightarrow> (\<forall>z. p(z,x) \<longleftrightarrow> p(z,y)))
- \<longrightarrow> (\<forall>x. \<forall>y. q(x,y) \<longleftrightarrow> q(y,x))"
+ \<open>(\<forall>x. \<forall>y. q(x,y) \<longleftrightarrow> (\<forall>z. p(z,x) \<longleftrightarrow> p(z,y)))
+ \<longrightarrow> (\<forall>x. \<forall>y. q(x,y) \<longleftrightarrow> q(y,x))\<close>
by blast
text \<open>
@@ -328,69 +328,69 @@
text\<open>44\<close>
lemma
- "(\<forall>x. f(x) \<longrightarrow> (\<exists>y. g(y) \<and> h(x,y) \<and> (\<exists>y. g(y) \<and> \<not> h(x,y)))) \<and>
+ \<open>(\<forall>x. f(x) \<longrightarrow> (\<exists>y. g(y) \<and> h(x,y) \<and> (\<exists>y. g(y) \<and> \<not> h(x,y)))) \<and>
(\<exists>x. j(x) \<and> (\<forall>y. g(y) \<longrightarrow> h(x,y)))
- \<longrightarrow> (\<exists>x. j(x) \<and> \<not> f(x))"
+ \<longrightarrow> (\<exists>x. j(x) \<and> \<not> f(x))\<close>
by blast
text\<open>45\<close>
lemma
- "(\<forall>x. f(x) \<and> (\<forall>y. g(y) \<and> h(x,y) \<longrightarrow> j(x,y))
+ \<open>(\<forall>x. f(x) \<and> (\<forall>y. g(y) \<and> h(x,y) \<longrightarrow> j(x,y))
\<longrightarrow> (\<forall>y. g(y) \<and> h(x,y) \<longrightarrow> k(y))) \<and>
\<not> (\<exists>y. l(y) \<and> k(y)) \<and>
(\<exists>x. f(x) \<and> (\<forall>y. h(x,y) \<longrightarrow> l(y)) \<and> (\<forall>y. g(y) \<and> h(x,y) \<longrightarrow> j(x,y)))
- \<longrightarrow> (\<exists>x. f(x) \<and> \<not> (\<exists>y. g(y) \<and> h(x,y)))"
+ \<longrightarrow> (\<exists>x. f(x) \<and> \<not> (\<exists>y. g(y) \<and> h(x,y)))\<close>
by blast
text\<open>46\<close>
lemma
- "(\<forall>x. f(x) \<and> (\<forall>y. f(y) \<and> h(y,x) \<longrightarrow> g(y)) \<longrightarrow> g(x)) \<and>
+ \<open>(\<forall>x. f(x) \<and> (\<forall>y. f(y) \<and> h(y,x) \<longrightarrow> g(y)) \<longrightarrow> g(x)) \<and>
((\<exists>x. f(x) \<and> \<not> g(x)) \<longrightarrow>
(\<exists>x. f(x) \<and> \<not> g(x) \<and> (\<forall>y. f(y) \<and> \<not> g(y) \<longrightarrow> j(x,y)))) \<and>
(\<forall>x y. f(x) \<and> f(y) \<and> h(x,y) \<longrightarrow> \<not> j(y,x))
- \<longrightarrow> (\<forall>x. f(x) \<longrightarrow> g(x))"
+ \<longrightarrow> (\<forall>x. f(x) \<longrightarrow> g(x))\<close>
by blast
subsection \<open>Problems (mainly) involving equality or functions\<close>
text\<open>48\<close>
-lemma "(a = b \<or> c = d) \<and> (a = c \<or> b = d) \<longrightarrow> a = d \<or> b = c"
+lemma \<open>(a = b \<or> c = d) \<and> (a = c \<or> b = d) \<longrightarrow> a = d \<or> b = c\<close>
by blast
text\<open>49. NOT PROVED AUTOMATICALLY. Hard because it involves substitution for
Vars; the type constraint ensures that x,y,z have the same type as a,b,u.\<close>
lemma
- "(\<exists>x y::'a. \<forall>z. z = x \<or> z = y) \<and> P(a) \<and> P(b) \<and> a \<noteq> b \<longrightarrow> (\<forall>u::'a. P(u))"
+ \<open>(\<exists>x y::'a. \<forall>z. z = x \<or> z = y) \<and> P(a) \<and> P(b) \<and> a \<noteq> b \<longrightarrow> (\<forall>u::'a. P(u))\<close>
apply safe
- apply (rule_tac x = a in allE, assumption)
- apply (rule_tac x = b in allE, assumption)
+ apply (rule_tac x = \<open>a\<close> in allE, assumption)
+ apply (rule_tac x = \<open>b\<close> in allE, assumption)
apply fast \<comment> \<open>blast's treatment of equality can't do it\<close>
done
text\<open>50. (What has this to do with equality?)\<close>
-lemma "(\<forall>x. P(a,x) \<or> (\<forall>y. P(x,y))) \<longrightarrow> (\<exists>x. \<forall>y. P(x,y))"
+lemma \<open>(\<forall>x. P(a,x) \<or> (\<forall>y. P(x,y))) \<longrightarrow> (\<exists>x. \<forall>y. P(x,y))\<close>
by blast
text\<open>51\<close>
lemma
- "(\<exists>z w. \<forall>x y. P(x,y) \<longleftrightarrow> (x = z \<and> y = w)) \<longrightarrow>
- (\<exists>z. \<forall>x. \<exists>w. (\<forall>y. P(x,y) \<longleftrightarrow> y=w) \<longleftrightarrow> x = z)"
+ \<open>(\<exists>z w. \<forall>x y. P(x,y) \<longleftrightarrow> (x = z \<and> y = w)) \<longrightarrow>
+ (\<exists>z. \<forall>x. \<exists>w. (\<forall>y. P(x,y) \<longleftrightarrow> y=w) \<longleftrightarrow> x = z)\<close>
by blast
text\<open>52\<close>
text\<open>Almost the same as 51.\<close>
lemma
- "(\<exists>z w. \<forall>x y. P(x,y) \<longleftrightarrow> (x = z \<and> y = w)) \<longrightarrow>
- (\<exists>w. \<forall>y. \<exists>z. (\<forall>x. P(x,y) \<longleftrightarrow> x = z) \<longleftrightarrow> y = w)"
+ \<open>(\<exists>z w. \<forall>x y. P(x,y) \<longleftrightarrow> (x = z \<and> y = w)) \<longrightarrow>
+ (\<exists>w. \<forall>y. \<exists>z. (\<forall>x. P(x,y) \<longleftrightarrow> x = z) \<longleftrightarrow> y = w)\<close>
by blast
text\<open>55\<close>
text\<open>Non-equational version, from Manthey and Bry, CADE-9 (Springer, 1988).
fast DISCOVERS who killed Agatha.\<close>
schematic_goal
- "lives(agatha) \<and> lives(butler) \<and> lives(charles) \<and>
+ \<open>lives(agatha) \<and> lives(butler) \<and> lives(charles) \<and>
(killed(agatha,agatha) \<or> killed(butler,agatha) \<or> killed(charles,agatha)) \<and>
(\<forall>x y. killed(x,y) \<longrightarrow> hates(x,y) \<and> \<not> richer(x,y)) \<and>
(\<forall>x. hates(agatha,x) \<longrightarrow> \<not> hates(charles,x)) \<and>
@@ -398,60 +398,60 @@
(\<forall>x. lives(x) \<and> \<not> richer(x,agatha) \<longrightarrow> hates(butler,x)) \<and>
(\<forall>x. hates(agatha,x) \<longrightarrow> hates(butler,x)) \<and>
(\<forall>x. \<not> hates(x,agatha) \<or> \<not> hates(x,butler) \<or> \<not> hates(x,charles)) \<longrightarrow>
- killed(?who,agatha)"
+ killed(?who,agatha)\<close>
by fast \<comment> \<open>MUCH faster than blast\<close>
text\<open>56\<close>
-lemma "(\<forall>x. (\<exists>y. P(y) \<and> x = f(y)) \<longrightarrow> P(x)) \<longleftrightarrow> (\<forall>x. P(x) \<longrightarrow> P(f(x)))"
+lemma \<open>(\<forall>x. (\<exists>y. P(y) \<and> x = f(y)) \<longrightarrow> P(x)) \<longleftrightarrow> (\<forall>x. P(x) \<longrightarrow> P(f(x)))\<close>
by blast
text\<open>57\<close>
lemma
- "P(f(a,b), f(b,c)) \<and> P(f(b,c), f(a,c)) \<and>
- (\<forall>x y z. P(x,y) \<and> P(y,z) \<longrightarrow> P(x,z)) \<longrightarrow> P(f(a,b), f(a,c))"
+ \<open>P(f(a,b), f(b,c)) \<and> P(f(b,c), f(a,c)) \<and>
+ (\<forall>x y z. P(x,y) \<and> P(y,z) \<longrightarrow> P(x,z)) \<longrightarrow> P(f(a,b), f(a,c))\<close>
by blast
text\<open>58 NOT PROVED AUTOMATICALLY\<close>
-lemma "(\<forall>x y. f(x) = g(y)) \<longrightarrow> (\<forall>x y. f(f(x)) = f(g(y)))"
+lemma \<open>(\<forall>x y. f(x) = g(y)) \<longrightarrow> (\<forall>x y. f(f(x)) = f(g(y)))\<close>
by (slow elim: subst_context)
text\<open>59\<close>
-lemma "(\<forall>x. P(x) \<longleftrightarrow> \<not> P(f(x))) \<longrightarrow> (\<exists>x. P(x) \<and> \<not> P(f(x)))"
+lemma \<open>(\<forall>x. P(x) \<longleftrightarrow> \<not> P(f(x))) \<longrightarrow> (\<exists>x. P(x) \<and> \<not> P(f(x)))\<close>
by blast
text\<open>60\<close>
-lemma "\<forall>x. P(x,f(x)) \<longleftrightarrow> (\<exists>y. (\<forall>z. P(z,y) \<longrightarrow> P(z,f(x))) \<and> P(x,y))"
+lemma \<open>\<forall>x. P(x,f(x)) \<longleftrightarrow> (\<exists>y. (\<forall>z. P(z,y) \<longrightarrow> P(z,f(x))) \<and> P(x,y))\<close>
by blast
text\<open>62 as corrected in JAR 18 (1997), page 135\<close>
lemma
- "(\<forall>x. p(a) \<and> (p(x) \<longrightarrow> p(f(x))) \<longrightarrow> p(f(f(x)))) \<longleftrightarrow>
+ \<open>(\<forall>x. p(a) \<and> (p(x) \<longrightarrow> p(f(x))) \<longrightarrow> p(f(f(x)))) \<longleftrightarrow>
(\<forall>x. (\<not> p(a) \<or> p(x) \<or> p(f(f(x)))) \<and>
- (\<not> p(a) \<or> \<not> p(f(x)) \<or> p(f(f(x)))))"
+ (\<not> p(a) \<or> \<not> p(f(x)) \<or> p(f(f(x)))))\<close>
by blast
text \<open>From Davis, Obvious Logical Inferences, IJCAI-81, 530-531
fast indeed copes!\<close>
lemma
- "(\<forall>x. F(x) \<and> \<not> G(x) \<longrightarrow> (\<exists>y. H(x,y) \<and> J(y))) \<and>
+ \<open>(\<forall>x. F(x) \<and> \<not> G(x) \<longrightarrow> (\<exists>y. H(x,y) \<and> J(y))) \<and>
(\<exists>x. K(x) \<and> F(x) \<and> (\<forall>y. H(x,y) \<longrightarrow> K(y))) \<and>
- (\<forall>x. K(x) \<longrightarrow> \<not> G(x)) \<longrightarrow> (\<exists>x. K(x) \<and> J(x))"
+ (\<forall>x. K(x) \<longrightarrow> \<not> G(x)) \<longrightarrow> (\<exists>x. K(x) \<and> J(x))\<close>
by fast
text \<open>From Rudnicki, Obvious Inferences, JAR 3 (1987), 383-393.
It does seem obvious!\<close>
lemma
- "(\<forall>x. F(x) \<and> \<not> G(x) \<longrightarrow> (\<exists>y. H(x,y) \<and> J(y))) \<and>
+ \<open>(\<forall>x. F(x) \<and> \<not> G(x) \<longrightarrow> (\<exists>y. H(x,y) \<and> J(y))) \<and>
(\<exists>x. K(x) \<and> F(x) \<and> (\<forall>y. H(x,y) \<longrightarrow> K(y))) \<and>
- (\<forall>x. K(x) \<longrightarrow> \<not> G(x)) \<longrightarrow> (\<exists>x. K(x) \<longrightarrow> \<not> G(x))"
+ (\<forall>x. K(x) \<longrightarrow> \<not> G(x)) \<longrightarrow> (\<exists>x. K(x) \<longrightarrow> \<not> G(x))\<close>
by fast
text \<open>Halting problem: Formulation of Li Dafa (AAR Newsletter 27, Oct 1994.)
author U. Egly.\<close>
lemma
- "((\<exists>x. A(x) \<and> (\<forall>y. C(y) \<longrightarrow> (\<forall>z. D(x,y,z)))) \<longrightarrow>
+ \<open>((\<exists>x. A(x) \<and> (\<forall>y. C(y) \<longrightarrow> (\<forall>z. D(x,y,z)))) \<longrightarrow>
(\<exists>w. C(w) \<and> (\<forall>y. C(y) \<longrightarrow> (\<forall>z. D(w,y,z)))))
\<and>
(\<forall>w. C(w) \<and> (\<forall>u. C(u) \<longrightarrow> (\<forall>v. D(w,u,v))) \<longrightarrow>
@@ -466,7 +466,7 @@
(\<exists>v. C(v) \<and>
(\<forall>y. ((C(y) \<and> Q(w,y,y)) \<and> OO(w,g) \<longrightarrow> \<not> P(v,y)) \<and>
((C(y) \<and> Q(w,y,y)) \<and> OO(w,b) \<longrightarrow> P(v,y) \<and> OO(v,b)))))
- \<longrightarrow> \<not> (\<exists>x. A(x) \<and> (\<forall>y. C(y) \<longrightarrow> (\<forall>z. D(x,y,z))))"
+ \<longrightarrow> \<not> (\<exists>x. A(x) \<and> (\<forall>y. C(y) \<longrightarrow> (\<forall>z. D(x,y,z))))\<close>
by (blast 12)
\<comment> \<open>Needed because the search for depths below 12 is very slow.\<close>
@@ -476,7 +476,7 @@
p. 105.
\<close>
lemma
- "((\<exists>x. A(x) \<and> (\<forall>y. C(y) \<longrightarrow> (\<forall>z. D(x,y,z)))) \<longrightarrow>
+ \<open>((\<exists>x. A(x) \<and> (\<forall>y. C(y) \<longrightarrow> (\<forall>z. D(x,y,z)))) \<longrightarrow>
(\<exists>w. C(w) \<and> (\<forall>y. C(y) \<longrightarrow> (\<forall>z. D(w,y,z)))))
\<and>
(\<forall>w. C(w) \<and> (\<forall>u. C(u) \<longrightarrow> (\<forall>v. D(w,u,v))) \<longrightarrow>
@@ -495,23 +495,23 @@
\<longrightarrow>
(\<exists>u. C(u) \<and> (\<forall>y. (C(y) \<and> P(y,y) \<longrightarrow> \<not> P(u,y)) \<and>
(C(y) \<and> \<not> P(y,y) \<longrightarrow> P(u,y) \<and> OO(u,b)))))
- \<longrightarrow> \<not> (\<exists>x. A(x) \<and> (\<forall>y. C(y) \<longrightarrow> (\<forall>z. D(x,y,z))))"
+ \<longrightarrow> \<not> (\<exists>x. A(x) \<and> (\<forall>y. C(y) \<longrightarrow> (\<forall>z. D(x,y,z))))\<close>
by blast
text \<open>Challenge found on info-hol.\<close>
-lemma "\<forall>x. \<exists>v w. \<forall>y z. P(x) \<and> Q(y) \<longrightarrow> (P(v) \<or> R(w)) \<and> (R(z) \<longrightarrow> Q(v))"
+lemma \<open>\<forall>x. \<exists>v w. \<forall>y z. P(x) \<and> Q(y) \<longrightarrow> (P(v) \<or> R(w)) \<and> (R(z) \<longrightarrow> Q(v))\<close>
by blast
text \<open>
Attributed to Lewis Carroll by S. G. Pulman. The first or last assumption
can be deleted.\<close>
lemma
- "(\<forall>x. honest(x) \<and> industrious(x) \<longrightarrow> healthy(x)) \<and>
+ \<open>(\<forall>x. honest(x) \<and> industrious(x) \<longrightarrow> healthy(x)) \<and>
\<not> (\<exists>x. grocer(x) \<and> healthy(x)) \<and>
(\<forall>x. industrious(x) \<and> grocer(x) \<longrightarrow> honest(x)) \<and>
(\<forall>x. cyclist(x) \<longrightarrow> industrious(x)) \<and>
(\<forall>x. \<not> healthy(x) \<and> cyclist(x) \<longrightarrow> \<not> honest(x))
- \<longrightarrow> (\<forall>x. grocer(x) \<longrightarrow> \<not> cyclist(x))"
+ \<longrightarrow> (\<forall>x. grocer(x) \<longrightarrow> \<not> cyclist(x))\<close>
by blast
--- a/src/FOL/ex/Foundation.thy Thu Jan 03 21:48:05 2019 +0100
+++ b/src/FOL/ex/Foundation.thy Thu Jan 03 22:19:19 2019 +0100
@@ -9,7 +9,7 @@
imports IFOL
begin
-lemma "A \<and> B \<longrightarrow> (C \<longrightarrow> A \<and> C)"
+lemma \<open>A \<and> B \<longrightarrow> (C \<longrightarrow> A \<and> C)\<close>
apply (rule impI)
apply (rule impI)
apply (rule conjI)
@@ -20,9 +20,9 @@
text \<open>A form of conj-elimination\<close>
lemma
- assumes "A \<and> B"
- and "A \<Longrightarrow> B \<Longrightarrow> C"
- shows C
+ assumes \<open>A \<and> B\<close>
+ and \<open>A \<Longrightarrow> B \<Longrightarrow> C\<close>
+ shows \<open>C\<close>
apply (rule assms)
apply (rule conjunct1)
apply (rule assms)
@@ -31,26 +31,26 @@
done
lemma
- assumes "\<And>A. \<not> \<not> A \<Longrightarrow> A"
- shows "B \<or> \<not> B"
+ assumes \<open>\<And>A. \<not> \<not> A \<Longrightarrow> A\<close>
+ shows \<open>B \<or> \<not> B\<close>
apply (rule assms)
apply (rule notI)
-apply (rule_tac P = "\<not> B" in notE)
+apply (rule_tac P = \<open>\<not> B\<close> in notE)
apply (rule_tac [2] notI)
-apply (rule_tac [2] P = "B \<or> \<not> B" in notE)
+apply (rule_tac [2] P = \<open>B \<or> \<not> B\<close> in notE)
prefer 2 apply assumption
apply (rule_tac [2] disjI1)
prefer 2 apply assumption
apply (rule notI)
-apply (rule_tac P = "B \<or> \<not> B" in notE)
+apply (rule_tac P = \<open>B \<or> \<not> B\<close> in notE)
apply assumption
apply (rule disjI2)
apply assumption
done
lemma
- assumes "\<And>A. \<not> \<not> A \<Longrightarrow> A"
- shows "B \<or> \<not> B"
+ assumes \<open>\<And>A. \<not> \<not> A \<Longrightarrow> A\<close>
+ shows \<open>B \<or> \<not> B\<close>
apply (rule assms)
apply (rule notI)
apply (rule notE)
@@ -64,14 +64,14 @@
lemma
- assumes "A \<or> \<not> A"
- and "\<not> \<not> A"
- shows A
+ assumes \<open>A \<or> \<not> A\<close>
+ and \<open>\<not> \<not> A\<close>
+ shows \<open>A\<close>
apply (rule disjE)
apply (rule assms)
apply assumption
apply (rule FalseE)
-apply (rule_tac P = "\<not> A" in notE)
+apply (rule_tac P = \<open>\<not> A\<close> in notE)
apply (rule assms)
apply assumption
done
@@ -80,27 +80,27 @@
subsection "Examples with quantifiers"
lemma
- assumes "\<forall>z. G(z)"
- shows "\<forall>z. G(z) \<or> H(z)"
+ assumes \<open>\<forall>z. G(z)\<close>
+ shows \<open>\<forall>z. G(z) \<or> H(z)\<close>
apply (rule allI)
apply (rule disjI1)
apply (rule assms [THEN spec])
done
-lemma "\<forall>x. \<exists>y. x = y"
+lemma \<open>\<forall>x. \<exists>y. x = y\<close>
apply (rule allI)
apply (rule exI)
apply (rule refl)
done
-lemma "\<exists>y. \<forall>x. x = y"
+lemma \<open>\<exists>y. \<forall>x. x = y\<close>
apply (rule exI)
apply (rule allI)
apply (rule refl)?
oops
text \<open>Parallel lifting example.\<close>
-lemma "\<exists>u. \<forall>x. \<exists>v. \<forall>y. \<exists>w. P(u,x,v,y,w)"
+lemma \<open>\<exists>u. \<forall>x. \<exists>v. \<forall>y. \<exists>w. P(u,x,v,y,w)\<close>
apply (rule exI allI)
apply (rule exI allI)
apply (rule exI allI)
@@ -109,8 +109,8 @@
oops
lemma
- assumes "(\<exists>z. F(z)) \<and> B"
- shows "\<exists>z. F(z) \<and> B"
+ assumes \<open>(\<exists>z. F(z)) \<and> B\<close>
+ shows \<open>\<exists>z. F(z) \<and> B\<close>
apply (rule conjE)
apply (rule assms)
apply (rule exE)
@@ -122,7 +122,7 @@
done
text \<open>A bigger demonstration of quantifiers -- not in the paper.\<close>
-lemma "(\<exists>y. \<forall>x. Q(x,y)) \<longrightarrow> (\<forall>x. \<exists>y. Q(x,y))"
+lemma \<open>(\<exists>y. \<forall>x. Q(x,y)) \<longrightarrow> (\<forall>x. \<exists>y. Q(x,y))\<close>
apply (rule impI)
apply (rule allI)
apply (rule exE, assumption)
--- a/src/FOL/ex/If.thy Thu Jan 03 21:48:05 2019 +0100
+++ b/src/FOL/ex/If.thy Thu Jan 03 22:19:19 2019 +0100
@@ -9,16 +9,16 @@
imports FOL
begin
-definition "if" :: "[o,o,o]=>o"
- where "if(P,Q,R) \<equiv> P \<and> Q \<or> \<not> P \<and> R"
+definition "if" :: \<open>[o,o,o]=>o\<close>
+ where \<open>if(P,Q,R) \<equiv> P \<and> Q \<or> \<not> P \<and> R\<close>
-lemma ifI: "\<lbrakk>P \<Longrightarrow> Q; \<not> P \<Longrightarrow> R\<rbrakk> \<Longrightarrow> if(P,Q,R)"
+lemma ifI: \<open>\<lbrakk>P \<Longrightarrow> Q; \<not> P \<Longrightarrow> R\<rbrakk> \<Longrightarrow> if(P,Q,R)\<close>
unfolding if_def by blast
-lemma ifE: "\<lbrakk>if(P,Q,R); \<lbrakk>P; Q\<rbrakk> \<Longrightarrow> S; \<lbrakk>\<not> P; R\<rbrakk> \<Longrightarrow> S\<rbrakk> \<Longrightarrow> S"
+lemma ifE: \<open>\<lbrakk>if(P,Q,R); \<lbrakk>P; Q\<rbrakk> \<Longrightarrow> S; \<lbrakk>\<not> P; R\<rbrakk> \<Longrightarrow> S\<rbrakk> \<Longrightarrow> S\<close>
unfolding if_def by blast
-lemma if_commute: "if(P, if(Q,A,B), if(Q,C,D)) \<longleftrightarrow> if(Q, if(P,A,C), if(P,B,D))"
+lemma if_commute: \<open>if(P, if(Q,A,B), if(Q,C,D)) \<longleftrightarrow> if(Q, if(P,A,C), if(P,B,D))\<close>
apply (rule iffI)
apply (erule ifE)
apply (erule ifE)
@@ -30,27 +30,27 @@
declare ifI [intro!]
declare ifE [elim!]
-lemma if_commute: "if(P, if(Q,A,B), if(Q,C,D)) \<longleftrightarrow> if(Q, if(P,A,C), if(P,B,D))"
+lemma if_commute: \<open>if(P, if(Q,A,B), if(Q,C,D)) \<longleftrightarrow> if(Q, if(P,A,C), if(P,B,D))\<close>
by blast
-lemma "if(if(P,Q,R), A, B) \<longleftrightarrow> if(P, if(Q,A,B), if(R,A,B))"
+lemma \<open>if(if(P,Q,R), A, B) \<longleftrightarrow> if(P, if(Q,A,B), if(R,A,B))\<close>
by blast
text\<open>Trying again from the beginning in order to prove from the definitions\<close>
-lemma "if(if(P,Q,R), A, B) \<longleftrightarrow> if(P, if(Q,A,B), if(R,A,B))"
+lemma \<open>if(if(P,Q,R), A, B) \<longleftrightarrow> if(P, if(Q,A,B), if(R,A,B))\<close>
unfolding if_def by blast
text \<open>An invalid formula. High-level rules permit a simpler diagnosis.\<close>
-lemma "if(if(P,Q,R), A, B) \<longleftrightarrow> if(P, if(Q,A,B), if(R,B,A))"
+lemma \<open>if(if(P,Q,R), A, B) \<longleftrightarrow> if(P, if(Q,A,B), if(R,B,A))\<close>
apply auto
\<comment> \<open>The next step will fail unless subgoals remain\<close>
apply (tactic all_tac)
oops
text \<open>Trying again from the beginning in order to prove from the definitions.\<close>
-lemma "if(if(P,Q,R), A, B) \<longleftrightarrow> if(P, if(Q,A,B), if(R,B,A))"
+lemma \<open>if(if(P,Q,R), A, B) \<longleftrightarrow> if(P, if(Q,A,B), if(R,B,A))\<close>
unfolding if_def
apply auto
\<comment> \<open>The next step will fail unless subgoals remain\<close>
--- a/src/FOL/ex/Intro.thy Thu Jan 03 21:48:05 2019 +0100
+++ b/src/FOL/ex/Intro.thy Thu Jan 03 22:19:19 2019 +0100
@@ -13,7 +13,7 @@
subsubsection \<open>Some simple backward proofs\<close>
-lemma mythm: "P \<or> P \<longrightarrow> P"
+lemma mythm: \<open>P \<or> P \<longrightarrow> P\<close>
apply (rule impI)
apply (rule disjE)
prefer 3 apply (assumption)
@@ -21,7 +21,7 @@
apply assumption
done
-lemma "(P \<and> Q) \<or> R \<longrightarrow> (P \<or> R)"
+lemma \<open>(P \<and> Q) \<or> R \<longrightarrow> (P \<or> R)\<close>
apply (rule impI)
apply (erule disjE)
apply (drule conjunct1)
@@ -31,7 +31,7 @@
done
text \<open>Correct version, delaying use of \<open>spec\<close> until last.\<close>
-lemma "(\<forall>x y. P(x,y)) \<longrightarrow> (\<forall>z w. P(w,z))"
+lemma \<open>(\<forall>x y. P(x,y)) \<longrightarrow> (\<forall>z w. P(w,z))\<close>
apply (rule impI)
apply (rule allI)
apply (rule allI)
@@ -43,12 +43,12 @@
subsubsection \<open>Demonstration of \<open>fast\<close>\<close>
-lemma "(\<exists>y. \<forall>x. J(y,x) \<longleftrightarrow> \<not> J(x,x)) \<longrightarrow> \<not> (\<forall>x. \<exists>y. \<forall>z. J(z,y) \<longleftrightarrow> \<not> J(z,x))"
+lemma \<open>(\<exists>y. \<forall>x. J(y,x) \<longleftrightarrow> \<not> J(x,x)) \<longrightarrow> \<not> (\<forall>x. \<exists>y. \<forall>z. J(z,y) \<longleftrightarrow> \<not> J(z,x))\<close>
apply fast
done
-lemma "\<forall>x. P(x,f(x)) \<longleftrightarrow> (\<exists>y. (\<forall>z. P(z,y) \<longrightarrow> P(z,f(x))) \<and> P(x,y))"
+lemma \<open>\<forall>x. P(x,f(x)) \<longleftrightarrow> (\<exists>y. (\<forall>z. P(z,y) \<longrightarrow> P(z,f(x))) \<and> P(x,y))\<close>
apply fast
done
@@ -56,9 +56,9 @@
subsubsection \<open>Derivation of conjunction elimination rule\<close>
lemma
- assumes major: "P \<and> Q"
- and minor: "\<lbrakk>P; Q\<rbrakk> \<Longrightarrow> R"
- shows R
+ assumes major: \<open>P \<and> Q\<close>
+ and minor: \<open>\<lbrakk>P; Q\<rbrakk> \<Longrightarrow> R\<close>
+ shows \<open>R\<close>
apply (rule minor)
apply (rule major [THEN conjunct1])
apply (rule major [THEN conjunct2])
@@ -70,8 +70,8 @@
text \<open>Derivation of negation introduction\<close>
lemma
- assumes "P \<Longrightarrow> False"
- shows "\<not> P"
+ assumes \<open>P \<Longrightarrow> False\<close>
+ shows \<open>\<not> P\<close>
apply (unfold not_def)
apply (rule impI)
apply (rule assms)
@@ -79,9 +79,9 @@
done
lemma
- assumes major: "\<not> P"
- and minor: P
- shows R
+ assumes major: \<open>\<not> P\<close>
+ and minor: \<open>P\<close>
+ shows \<open>R\<close>
apply (rule FalseE)
apply (rule mp)
apply (rule major [unfolded not_def])
@@ -90,9 +90,9 @@
text \<open>Alternative proof of the result above\<close>
lemma
- assumes major: "\<not> P"
- and minor: P
- shows R
+ assumes major: \<open>\<not> P\<close>
+ and minor: \<open>P\<close>
+ shows \<open>R\<close>
apply (rule minor [THEN major [unfolded not_def, THEN mp, THEN FalseE]])
done
--- a/src/FOL/ex/Intuitionistic.thy Thu Jan 03 21:48:05 2019 +0100
+++ b/src/FOL/ex/Intuitionistic.thy Thu Jan 03 22:19:19 2019 +0100
@@ -36,56 +36,56 @@
intuitionistically provable. Finally, if $P$ is a negation then $\neg\neg P$
is intuitionstically equivalent to $P$. [Andy Pitts]\<close>
-lemma "\<not> \<not> (P \<and> Q) \<longleftrightarrow> \<not> \<not> P \<and> \<not> \<not> Q"
+lemma \<open>\<not> \<not> (P \<and> Q) \<longleftrightarrow> \<not> \<not> P \<and> \<not> \<not> Q\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
-lemma "\<not> \<not> ((\<not> P \<longrightarrow> Q) \<longrightarrow> (\<not> P \<longrightarrow> \<not> Q) \<longrightarrow> P)"
+lemma \<open>\<not> \<not> ((\<not> P \<longrightarrow> Q) \<longrightarrow> (\<not> P \<longrightarrow> \<not> Q) \<longrightarrow> P)\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text \<open>Double-negation does NOT distribute over disjunction.\<close>
-lemma "\<not> \<not> (P \<longrightarrow> Q) \<longleftrightarrow> (\<not> \<not> P \<longrightarrow> \<not> \<not> Q)"
+lemma \<open>\<not> \<not> (P \<longrightarrow> Q) \<longleftrightarrow> (\<not> \<not> P \<longrightarrow> \<not> \<not> Q)\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
-lemma "\<not> \<not> \<not> P \<longleftrightarrow> \<not> P"
+lemma \<open>\<not> \<not> \<not> P \<longleftrightarrow> \<not> P\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
-lemma "\<not> \<not> ((P \<longrightarrow> Q \<or> R) \<longrightarrow> (P \<longrightarrow> Q) \<or> (P \<longrightarrow> R))"
+lemma \<open>\<not> \<not> ((P \<longrightarrow> Q \<or> R) \<longrightarrow> (P \<longrightarrow> Q) \<or> (P \<longrightarrow> R))\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
-lemma "(P \<longleftrightarrow> Q) \<longleftrightarrow> (Q \<longleftrightarrow> P)"
+lemma \<open>(P \<longleftrightarrow> Q) \<longleftrightarrow> (Q \<longleftrightarrow> P)\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
-lemma "((P \<longrightarrow> (Q \<or> (Q \<longrightarrow> R))) \<longrightarrow> R) \<longrightarrow> R"
+lemma \<open>((P \<longrightarrow> (Q \<or> (Q \<longrightarrow> R))) \<longrightarrow> R) \<longrightarrow> R\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
lemma
- "(((G \<longrightarrow> A) \<longrightarrow> J) \<longrightarrow> D \<longrightarrow> E) \<longrightarrow> (((H \<longrightarrow> B) \<longrightarrow> I) \<longrightarrow> C \<longrightarrow> J)
+ \<open>(((G \<longrightarrow> A) \<longrightarrow> J) \<longrightarrow> D \<longrightarrow> E) \<longrightarrow> (((H \<longrightarrow> B) \<longrightarrow> I) \<longrightarrow> C \<longrightarrow> J)
\<longrightarrow> (A \<longrightarrow> H) \<longrightarrow> F \<longrightarrow> G \<longrightarrow> (((C \<longrightarrow> B) \<longrightarrow> I) \<longrightarrow> D) \<longrightarrow> (A \<longrightarrow> C)
- \<longrightarrow> (((F \<longrightarrow> A) \<longrightarrow> B) \<longrightarrow> I) \<longrightarrow> E"
+ \<longrightarrow> (((F \<longrightarrow> A) \<longrightarrow> B) \<longrightarrow> I) \<longrightarrow> E\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
subsection \<open>Lemmas for the propositional double-negation translation\<close>
-lemma "P \<longrightarrow> \<not> \<not> P"
+lemma \<open>P \<longrightarrow> \<not> \<not> P\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
-lemma "\<not> \<not> (\<not> \<not> P \<longrightarrow> P)"
+lemma \<open>\<not> \<not> (\<not> \<not> P \<longrightarrow> P)\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
-lemma "\<not> \<not> P \<and> \<not> \<not> (P \<longrightarrow> Q) \<longrightarrow> \<not> \<not> Q"
+lemma \<open>\<not> \<not> P \<and> \<not> \<not> (P \<longrightarrow> Q) \<longrightarrow> \<not> \<not> Q\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text \<open>The following are classically but not constructively valid.
The attempt to prove them terminates quickly!\<close>
-lemma "((P \<longrightarrow> Q) \<longrightarrow> P) \<longrightarrow> P"
+lemma \<open>((P \<longrightarrow> Q) \<longrightarrow> P) \<longrightarrow> P\<close>
apply (tactic \<open>IntPr.fast_tac @{context} 1\<close>)?
apply (rule asm_rl) \<comment> \<open>Checks that subgoals remain: proof failed.\<close>
oops
-lemma "(P \<and> Q \<longrightarrow> R) \<longrightarrow> (P \<longrightarrow> R) \<or> (Q \<longrightarrow> R)"
+lemma \<open>(P \<and> Q \<longrightarrow> R) \<longrightarrow> (P \<longrightarrow> R) \<or> (Q \<longrightarrow> R)\<close>
apply (tactic \<open>IntPr.fast_tac @{context} 1\<close>)?
apply (rule asm_rl) \<comment> \<open>Checks that subgoals remain: proof failed.\<close>
oops
@@ -95,19 +95,19 @@
text \<open>de Bruijn formula with three predicates\<close>
lemma
- "((P \<longleftrightarrow> Q) \<longrightarrow> P \<and> Q \<and> R) \<and>
+ \<open>((P \<longleftrightarrow> Q) \<longrightarrow> P \<and> Q \<and> R) \<and>
((Q \<longleftrightarrow> R) \<longrightarrow> P \<and> Q \<and> R) \<and>
- ((R \<longleftrightarrow> P) \<longrightarrow> P \<and> Q \<and> R) \<longrightarrow> P \<and> Q \<and> R"
+ ((R \<longleftrightarrow> P) \<longrightarrow> P \<and> Q \<and> R) \<longrightarrow> P \<and> Q \<and> R\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text \<open>de Bruijn formula with five predicates\<close>
lemma
- "((P \<longleftrightarrow> Q) \<longrightarrow> P \<and> Q \<and> R \<and> S \<and> T) \<and>
+ \<open>((P \<longleftrightarrow> Q) \<longrightarrow> P \<and> Q \<and> R \<and> S \<and> T) \<and>
((Q \<longleftrightarrow> R) \<longrightarrow> P \<and> Q \<and> R \<and> S \<and> T) \<and>
((R \<longleftrightarrow> S) \<longrightarrow> P \<and> Q \<and> R \<and> S \<and> T) \<and>
((S \<longleftrightarrow> T) \<longrightarrow> P \<and> Q \<and> R \<and> S \<and> T) \<and>
- ((T \<longleftrightarrow> P) \<longrightarrow> P \<and> Q \<and> R \<and> S \<and> T) \<longrightarrow> P \<and> Q \<and> R \<and> S \<and> T"
+ ((T \<longleftrightarrow> P) \<longrightarrow> P \<and> Q \<and> R \<and> S \<and> T) \<longrightarrow> P \<and> Q \<and> R \<and> S \<and> T\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
@@ -119,98 +119,98 @@
text\<open>Problem 1.1\<close>
lemma
- "(\<forall>x. \<exists>y. \<forall>z. p(x) \<and> q(y) \<and> r(z)) \<longleftrightarrow>
- (\<forall>z. \<exists>y. \<forall>x. p(x) \<and> q(y) \<and> r(z))"
+ \<open>(\<forall>x. \<exists>y. \<forall>z. p(x) \<and> q(y) \<and> r(z)) \<longleftrightarrow>
+ (\<forall>z. \<exists>y. \<forall>x. p(x) \<and> q(y) \<and> r(z))\<close>
by (tactic \<open>IntPr.best_dup_tac @{context} 1\<close>) \<comment> \<open>SLOW\<close>
text\<open>Problem 3.1\<close>
-lemma "\<not> (\<exists>x. \<forall>y. mem(y,x) \<longleftrightarrow> \<not> mem(x,x))"
+lemma \<open>\<not> (\<exists>x. \<forall>y. mem(y,x) \<longleftrightarrow> \<not> mem(x,x))\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>Problem 4.1: hopeless!\<close>
lemma
- "(\<forall>x. p(x) \<longrightarrow> p(h(x)) \<or> p(g(x))) \<and> (\<exists>x. p(x)) \<and> (\<forall>x. \<not> p(h(x)))
- \<longrightarrow> (\<exists>x. p(g(g(g(g(g(x)))))))"
+ \<open>(\<forall>x. p(x) \<longrightarrow> p(h(x)) \<or> p(g(x))) \<and> (\<exists>x. p(x)) \<and> (\<forall>x. \<not> p(h(x)))
+ \<longrightarrow> (\<exists>x. p(g(g(g(g(g(x)))))))\<close>
oops
subsection \<open>Intuitionistic FOL: propositional problems based on Pelletier.\<close>
text\<open>\<open>\<not>\<not>\<close>1\<close>
-lemma "\<not> \<not> ((P \<longrightarrow> Q) \<longleftrightarrow> (\<not> Q \<longrightarrow> \<not> P))"
+lemma \<open>\<not> \<not> ((P \<longrightarrow> Q) \<longleftrightarrow> (\<not> Q \<longrightarrow> \<not> P))\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>\<open>\<not>\<not>\<close>2\<close>
-lemma "\<not> \<not> (\<not> \<not> P \<longleftrightarrow> P)"
+lemma \<open>\<not> \<not> (\<not> \<not> P \<longleftrightarrow> P)\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>3\<close>
-lemma "\<not> (P \<longrightarrow> Q) \<longrightarrow> (Q \<longrightarrow> P)"
+lemma \<open>\<not> (P \<longrightarrow> Q) \<longrightarrow> (Q \<longrightarrow> P)\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>\<open>\<not>\<not>\<close>4\<close>
-lemma "\<not> \<not> ((\<not> P \<longrightarrow> Q) \<longleftrightarrow> (\<not> Q \<longrightarrow> P))"
+lemma \<open>\<not> \<not> ((\<not> P \<longrightarrow> Q) \<longleftrightarrow> (\<not> Q \<longrightarrow> P))\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>\<open>\<not>\<not>\<close>5\<close>
-lemma "\<not> \<not> ((P \<or> Q \<longrightarrow> P \<or> R) \<longrightarrow> P \<or> (Q \<longrightarrow> R))"
+lemma \<open>\<not> \<not> ((P \<or> Q \<longrightarrow> P \<or> R) \<longrightarrow> P \<or> (Q \<longrightarrow> R))\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>\<open>\<not>\<not>\<close>6\<close>
-lemma "\<not> \<not> (P \<or> \<not> P)"
+lemma \<open>\<not> \<not> (P \<or> \<not> P)\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>\<open>\<not>\<not>\<close>7\<close>
-lemma "\<not> \<not> (P \<or> \<not> \<not> \<not> P)"
+lemma \<open>\<not> \<not> (P \<or> \<not> \<not> \<not> P)\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>\<open>\<not>\<not>\<close>8. Peirce's law\<close>
-lemma "\<not> \<not> (((P \<longrightarrow> Q) \<longrightarrow> P) \<longrightarrow> P)"
+lemma \<open>\<not> \<not> (((P \<longrightarrow> Q) \<longrightarrow> P) \<longrightarrow> P)\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>9\<close>
-lemma "((P \<or> Q) \<and> (\<not> P \<or> Q) \<and> (P \<or> \<not> Q)) \<longrightarrow> \<not> (\<not> P \<or> \<not> Q)"
+lemma \<open>((P \<or> Q) \<and> (\<not> P \<or> Q) \<and> (P \<or> \<not> Q)) \<longrightarrow> \<not> (\<not> P \<or> \<not> Q)\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>10\<close>
-lemma "(Q \<longrightarrow> R) \<longrightarrow> (R \<longrightarrow> P \<and> Q) \<longrightarrow> (P \<longrightarrow> (Q \<or> R)) \<longrightarrow> (P \<longleftrightarrow> Q)"
+lemma \<open>(Q \<longrightarrow> R) \<longrightarrow> (R \<longrightarrow> P \<and> Q) \<longrightarrow> (P \<longrightarrow> (Q \<or> R)) \<longrightarrow> (P \<longleftrightarrow> Q)\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
subsection\<open>11. Proved in each direction (incorrectly, says Pelletier!!)\<close>
-lemma "P \<longleftrightarrow> P"
+lemma \<open>P \<longleftrightarrow> P\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>\<open>\<not>\<not>\<close>12. Dijkstra's law\<close>
-lemma "\<not> \<not> (((P \<longleftrightarrow> Q) \<longleftrightarrow> R) \<longleftrightarrow> (P \<longleftrightarrow> (Q \<longleftrightarrow> R)))"
+lemma \<open>\<not> \<not> (((P \<longleftrightarrow> Q) \<longleftrightarrow> R) \<longleftrightarrow> (P \<longleftrightarrow> (Q \<longleftrightarrow> R)))\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
-lemma "((P \<longleftrightarrow> Q) \<longleftrightarrow> R) \<longrightarrow> \<not> \<not> (P \<longleftrightarrow> (Q \<longleftrightarrow> R))"
+lemma \<open>((P \<longleftrightarrow> Q) \<longleftrightarrow> R) \<longrightarrow> \<not> \<not> (P \<longleftrightarrow> (Q \<longleftrightarrow> R))\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>13. Distributive law\<close>
-lemma "P \<or> (Q \<and> R) \<longleftrightarrow> (P \<or> Q) \<and> (P \<or> R)"
+lemma \<open>P \<or> (Q \<and> R) \<longleftrightarrow> (P \<or> Q) \<and> (P \<or> R)\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>\<open>\<not>\<not>\<close>14\<close>
-lemma "\<not> \<not> ((P \<longleftrightarrow> Q) \<longleftrightarrow> ((Q \<or> \<not> P) \<and> (\<not> Q \<or> P)))"
+lemma \<open>\<not> \<not> ((P \<longleftrightarrow> Q) \<longleftrightarrow> ((Q \<or> \<not> P) \<and> (\<not> Q \<or> P)))\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>\<open>\<not>\<not>\<close>15\<close>
-lemma "\<not> \<not> ((P \<longrightarrow> Q) \<longleftrightarrow> (\<not> P \<or> Q))"
+lemma \<open>\<not> \<not> ((P \<longrightarrow> Q) \<longleftrightarrow> (\<not> P \<or> Q))\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>\<open>\<not>\<not>\<close>16\<close>
-lemma "\<not> \<not> ((P \<longrightarrow> Q) \<or> (Q \<longrightarrow> P))"
+lemma \<open>\<not> \<not> ((P \<longrightarrow> Q) \<or> (Q \<longrightarrow> P))\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>\<open>\<not>\<not>\<close>17\<close>
-lemma "\<not> \<not> (((P \<and> (Q \<longrightarrow> R)) \<longrightarrow> S) \<longleftrightarrow> ((\<not> P \<or> Q \<or> S) \<and> (\<not> P \<or> \<not> R \<or> S)))"
+lemma \<open>\<not> \<not> (((P \<and> (Q \<longrightarrow> R)) \<longrightarrow> S) \<longleftrightarrow> ((\<not> P \<or> Q \<or> S) \<and> (\<not> P \<or> \<not> R \<or> S)))\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text \<open>Dijkstra's ``Golden Rule''\<close>
-lemma "(P \<and> Q) \<longleftrightarrow> P \<longleftrightarrow> Q \<longleftrightarrow> (P \<or> Q)"
+lemma \<open>(P \<and> Q) \<longleftrightarrow> P \<longleftrightarrow> Q \<longleftrightarrow> (P \<or> Q)\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
@@ -218,47 +218,47 @@
subsection \<open>The converse is classical in the following implications \dots\<close>
-lemma "(\<exists>x. P(x) \<longrightarrow> Q) \<longrightarrow> (\<forall>x. P(x)) \<longrightarrow> Q"
+lemma \<open>(\<exists>x. P(x) \<longrightarrow> Q) \<longrightarrow> (\<forall>x. P(x)) \<longrightarrow> Q\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
-lemma "((\<forall>x. P(x)) \<longrightarrow> Q) \<longrightarrow> \<not> (\<forall>x. P(x) \<and> \<not> Q)"
+lemma \<open>((\<forall>x. P(x)) \<longrightarrow> Q) \<longrightarrow> \<not> (\<forall>x. P(x) \<and> \<not> Q)\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
-lemma "((\<forall>x. \<not> P(x)) \<longrightarrow> Q) \<longrightarrow> \<not> (\<forall>x. \<not> (P(x) \<or> Q))"
+lemma \<open>((\<forall>x. \<not> P(x)) \<longrightarrow> Q) \<longrightarrow> \<not> (\<forall>x. \<not> (P(x) \<or> Q))\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
-lemma "(\<forall>x. P(x)) \<or> Q \<longrightarrow> (\<forall>x. P(x) \<or> Q)"
+lemma \<open>(\<forall>x. P(x)) \<or> Q \<longrightarrow> (\<forall>x. P(x) \<or> Q)\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
-lemma "(\<exists>x. P \<longrightarrow> Q(x)) \<longrightarrow> (P \<longrightarrow> (\<exists>x. Q(x)))"
+lemma \<open>(\<exists>x. P \<longrightarrow> Q(x)) \<longrightarrow> (P \<longrightarrow> (\<exists>x. Q(x)))\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
subsection \<open>The following are not constructively valid!\<close>
text \<open>The attempt to prove them terminates quickly!\<close>
-lemma "((\<forall>x. P(x)) \<longrightarrow> Q) \<longrightarrow> (\<exists>x. P(x) \<longrightarrow> Q)"
+lemma \<open>((\<forall>x. P(x)) \<longrightarrow> Q) \<longrightarrow> (\<exists>x. P(x) \<longrightarrow> Q)\<close>
apply (tactic \<open>IntPr.fast_tac @{context} 1\<close>)?
apply (rule asm_rl) \<comment> \<open>Checks that subgoals remain: proof failed.\<close>
oops
-lemma "(P \<longrightarrow> (\<exists>x. Q(x))) \<longrightarrow> (\<exists>x. P \<longrightarrow> Q(x))"
+lemma \<open>(P \<longrightarrow> (\<exists>x. Q(x))) \<longrightarrow> (\<exists>x. P \<longrightarrow> Q(x))\<close>
apply (tactic \<open>IntPr.fast_tac @{context} 1\<close>)?
apply (rule asm_rl) \<comment> \<open>Checks that subgoals remain: proof failed.\<close>
oops
-lemma "(\<forall>x. P(x) \<or> Q) \<longrightarrow> ((\<forall>x. P(x)) \<or> Q)"
+lemma \<open>(\<forall>x. P(x) \<or> Q) \<longrightarrow> ((\<forall>x. P(x)) \<or> Q)\<close>
apply (tactic \<open>IntPr.fast_tac @{context} 1\<close>)?
apply (rule asm_rl) \<comment> \<open>Checks that subgoals remain: proof failed.\<close>
oops
-lemma "(\<forall>x. \<not> \<not> P(x)) \<longrightarrow> \<not> \<not> (\<forall>x. P(x))"
+lemma \<open>(\<forall>x. \<not> \<not> P(x)) \<longrightarrow> \<not> \<not> (\<forall>x. P(x))\<close>
apply (tactic \<open>IntPr.fast_tac @{context} 1\<close>)?
apply (rule asm_rl) \<comment> \<open>Checks that subgoals remain: proof failed.\<close>
oops
text \<open>Classically but not intuitionistically valid. Proved by a bug in 1986!\<close>
-lemma "\<exists>x. Q(x) \<longrightarrow> (\<forall>x. Q(x))"
+lemma \<open>\<exists>x. Q(x) \<longrightarrow> (\<forall>x. Q(x))\<close>
apply (tactic \<open>IntPr.fast_tac @{context} 1\<close>)?
apply (rule asm_rl) \<comment> \<open>Checks that subgoals remain: proof failed.\<close>
oops
@@ -272,36 +272,36 @@
\<close>
text\<open>\<open>\<not>\<not>\<close>18\<close>
-lemma "\<not> \<not> (\<exists>y. \<forall>x. P(y) \<longrightarrow> P(x))"
+lemma \<open>\<not> \<not> (\<exists>y. \<forall>x. P(y) \<longrightarrow> P(x))\<close>
oops \<comment> \<open>NOT PROVED\<close>
text\<open>\<open>\<not>\<not>\<close>19\<close>
-lemma "\<not> \<not> (\<exists>x. \<forall>y z. (P(y) \<longrightarrow> Q(z)) \<longrightarrow> (P(x) \<longrightarrow> Q(x)))"
+lemma \<open>\<not> \<not> (\<exists>x. \<forall>y z. (P(y) \<longrightarrow> Q(z)) \<longrightarrow> (P(x) \<longrightarrow> Q(x)))\<close>
oops \<comment> \<open>NOT PROVED\<close>
text\<open>20\<close>
lemma
- "(\<forall>x y. \<exists>z. \<forall>w. (P(x) \<and> Q(y) \<longrightarrow> R(z) \<and> S(w)))
- \<longrightarrow> (\<exists>x y. P(x) \<and> Q(y)) \<longrightarrow> (\<exists>z. R(z))"
+ \<open>(\<forall>x y. \<exists>z. \<forall>w. (P(x) \<and> Q(y) \<longrightarrow> R(z) \<and> S(w)))
+ \<longrightarrow> (\<exists>x y. P(x) \<and> Q(y)) \<longrightarrow> (\<exists>z. R(z))\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>21\<close>
-lemma "(\<exists>x. P \<longrightarrow> Q(x)) \<and> (\<exists>x. Q(x) \<longrightarrow> P) \<longrightarrow> \<not> \<not> (\<exists>x. P \<longleftrightarrow> Q(x))"
+lemma \<open>(\<exists>x. P \<longrightarrow> Q(x)) \<and> (\<exists>x. Q(x) \<longrightarrow> P) \<longrightarrow> \<not> \<not> (\<exists>x. P \<longleftrightarrow> Q(x))\<close>
oops \<comment> \<open>NOT PROVED; needs quantifier duplication\<close>
text\<open>22\<close>
-lemma "(\<forall>x. P \<longleftrightarrow> Q(x)) \<longrightarrow> (P \<longleftrightarrow> (\<forall>x. Q(x)))"
+lemma \<open>(\<forall>x. P \<longleftrightarrow> Q(x)) \<longrightarrow> (P \<longleftrightarrow> (\<forall>x. Q(x)))\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>\<open>\<not>\<not>\<close>23\<close>
-lemma "\<not> \<not> ((\<forall>x. P \<or> Q(x)) \<longleftrightarrow> (P \<or> (\<forall>x. Q(x))))"
+lemma \<open>\<not> \<not> ((\<forall>x. P \<or> Q(x)) \<longleftrightarrow> (P \<or> (\<forall>x. Q(x))))\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>24\<close>
lemma
- "\<not> (\<exists>x. S(x) \<and> Q(x)) \<and> (\<forall>x. P(x) \<longrightarrow> Q(x) \<or> R(x)) \<and>
+ \<open>\<not> (\<exists>x. S(x) \<and> Q(x)) \<and> (\<forall>x. P(x) \<longrightarrow> Q(x) \<or> R(x)) \<and>
(\<not> (\<exists>x. P(x)) \<longrightarrow> (\<exists>x. Q(x))) \<and> (\<forall>x. Q(x) \<or> R(x) \<longrightarrow> S(x))
- \<longrightarrow> \<not> \<not> (\<exists>x. P(x) \<and> R(x))"
+ \<longrightarrow> \<not> \<not> (\<exists>x. P(x) \<and> R(x))\<close>
text \<open>
Not clear why \<open>fast_tac\<close>, \<open>best_tac\<close>, \<open>ASTAR\<close> and
\<open>ITER_DEEPEN\<close> all take forever.
@@ -314,139 +314,139 @@
text\<open>25\<close>
lemma
- "(\<exists>x. P(x)) \<and>
+ \<open>(\<exists>x. P(x)) \<and>
(\<forall>x. L(x) \<longrightarrow> \<not> (M(x) \<and> R(x))) \<and>
(\<forall>x. P(x) \<longrightarrow> (M(x) \<and> L(x))) \<and>
((\<forall>x. P(x) \<longrightarrow> Q(x)) \<or> (\<exists>x. P(x) \<and> R(x)))
- \<longrightarrow> (\<exists>x. Q(x) \<and> P(x))"
+ \<longrightarrow> (\<exists>x. Q(x) \<and> P(x))\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>\<open>\<not>\<not>\<close>26\<close>
lemma
- "(\<not> \<not> (\<exists>x. p(x)) \<longleftrightarrow> \<not> \<not> (\<exists>x. q(x))) \<and>
+ \<open>(\<not> \<not> (\<exists>x. p(x)) \<longleftrightarrow> \<not> \<not> (\<exists>x. q(x))) \<and>
(\<forall>x. \<forall>y. p(x) \<and> q(y) \<longrightarrow> (r(x) \<longleftrightarrow> s(y)))
- \<longrightarrow> ((\<forall>x. p(x) \<longrightarrow> r(x)) \<longleftrightarrow> (\<forall>x. q(x) \<longrightarrow> s(x)))"
+ \<longrightarrow> ((\<forall>x. p(x) \<longrightarrow> r(x)) \<longleftrightarrow> (\<forall>x. q(x) \<longrightarrow> s(x)))\<close>
oops \<comment> \<open>NOT PROVED\<close>
text\<open>27\<close>
lemma
- "(\<exists>x. P(x) \<and> \<not> Q(x)) \<and>
+ \<open>(\<exists>x. P(x) \<and> \<not> Q(x)) \<and>
(\<forall>x. P(x) \<longrightarrow> R(x)) \<and>
(\<forall>x. M(x) \<and> L(x) \<longrightarrow> P(x)) \<and>
((\<exists>x. R(x) \<and> \<not> Q(x)) \<longrightarrow> (\<forall>x. L(x) \<longrightarrow> \<not> R(x)))
- \<longrightarrow> (\<forall>x. M(x) \<longrightarrow> \<not> L(x))"
+ \<longrightarrow> (\<forall>x. M(x) \<longrightarrow> \<not> L(x))\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>\<open>\<not>\<not>\<close>28. AMENDED\<close>
lemma
- "(\<forall>x. P(x) \<longrightarrow> (\<forall>x. Q(x))) \<and>
+ \<open>(\<forall>x. P(x) \<longrightarrow> (\<forall>x. Q(x))) \<and>
(\<not> \<not> (\<forall>x. Q(x) \<or> R(x)) \<longrightarrow> (\<exists>x. Q(x) \<and> S(x))) \<and>
(\<not> \<not> (\<exists>x. S(x)) \<longrightarrow> (\<forall>x. L(x) \<longrightarrow> M(x)))
- \<longrightarrow> (\<forall>x. P(x) \<and> L(x) \<longrightarrow> M(x))"
+ \<longrightarrow> (\<forall>x. P(x) \<and> L(x) \<longrightarrow> M(x))\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>29. Essentially the same as Principia Mathematica *11.71\<close>
lemma
- "(\<exists>x. P(x)) \<and> (\<exists>y. Q(y))
+ \<open>(\<exists>x. P(x)) \<and> (\<exists>y. Q(y))
\<longrightarrow> ((\<forall>x. P(x) \<longrightarrow> R(x)) \<and> (\<forall>y. Q(y) \<longrightarrow> S(y)) \<longleftrightarrow>
- (\<forall>x y. P(x) \<and> Q(y) \<longrightarrow> R(x) \<and> S(y)))"
+ (\<forall>x y. P(x) \<and> Q(y) \<longrightarrow> R(x) \<and> S(y)))\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>\<open>\<not>\<not>\<close>30\<close>
lemma
- "(\<forall>x. (P(x) \<or> Q(x)) \<longrightarrow> \<not> R(x)) \<and>
+ \<open>(\<forall>x. (P(x) \<or> Q(x)) \<longrightarrow> \<not> R(x)) \<and>
(\<forall>x. (Q(x) \<longrightarrow> \<not> S(x)) \<longrightarrow> P(x) \<and> R(x))
- \<longrightarrow> (\<forall>x. \<not> \<not> S(x))"
+ \<longrightarrow> (\<forall>x. \<not> \<not> S(x))\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>31\<close>
lemma
- "\<not> (\<exists>x. P(x) \<and> (Q(x) \<or> R(x))) \<and>
+ \<open>\<not> (\<exists>x. P(x) \<and> (Q(x) \<or> R(x))) \<and>
(\<exists>x. L(x) \<and> P(x)) \<and>
(\<forall>x. \<not> R(x) \<longrightarrow> M(x))
- \<longrightarrow> (\<exists>x. L(x) \<and> M(x))"
+ \<longrightarrow> (\<exists>x. L(x) \<and> M(x))\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>32\<close>
lemma
- "(\<forall>x. P(x) \<and> (Q(x) \<or> R(x)) \<longrightarrow> S(x)) \<and>
+ \<open>(\<forall>x. P(x) \<and> (Q(x) \<or> R(x)) \<longrightarrow> S(x)) \<and>
(\<forall>x. S(x) \<and> R(x) \<longrightarrow> L(x)) \<and>
(\<forall>x. M(x) \<longrightarrow> R(x))
- \<longrightarrow> (\<forall>x. P(x) \<and> M(x) \<longrightarrow> L(x))"
+ \<longrightarrow> (\<forall>x. P(x) \<and> M(x) \<longrightarrow> L(x))\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>\<open>\<not>\<not>\<close>33\<close>
lemma
- "(\<forall>x. \<not> \<not> (P(a) \<and> (P(x) \<longrightarrow> P(b)) \<longrightarrow> P(c))) \<longleftrightarrow>
- (\<forall>x. \<not> \<not> ((\<not> P(a) \<or> P(x) \<or> P(c)) \<and> (\<not> P(a) \<or> \<not> P(b) \<or> P(c))))"
+ \<open>(\<forall>x. \<not> \<not> (P(a) \<and> (P(x) \<longrightarrow> P(b)) \<longrightarrow> P(c))) \<longleftrightarrow>
+ (\<forall>x. \<not> \<not> ((\<not> P(a) \<or> P(x) \<or> P(c)) \<and> (\<not> P(a) \<or> \<not> P(b) \<or> P(c))))\<close>
apply (tactic \<open>IntPr.best_tac @{context} 1\<close>)
done
text\<open>36\<close>
lemma
- "(\<forall>x. \<exists>y. J(x,y)) \<and>
+ \<open>(\<forall>x. \<exists>y. J(x,y)) \<and>
(\<forall>x. \<exists>y. G(x,y)) \<and>
(\<forall>x y. J(x,y) \<or> G(x,y) \<longrightarrow> (\<forall>z. J(y,z) \<or> G(y,z) \<longrightarrow> H(x,z)))
- \<longrightarrow> (\<forall>x. \<exists>y. H(x,y))"
+ \<longrightarrow> (\<forall>x. \<exists>y. H(x,y))\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>37\<close>
lemma
- "(\<forall>z. \<exists>w. \<forall>x. \<exists>y.
+ \<open>(\<forall>z. \<exists>w. \<forall>x. \<exists>y.
\<not> \<not> (P(x,z) \<longrightarrow> P(y,w)) \<and> P(y,z) \<and> (P(y,w) \<longrightarrow> (\<exists>u. Q(u,w)))) \<and>
(\<forall>x z. \<not> P(x,z) \<longrightarrow> (\<exists>y. Q(y,z))) \<and>
(\<not> \<not> (\<exists>x y. Q(x,y)) \<longrightarrow> (\<forall>x. R(x,x)))
- \<longrightarrow> \<not> \<not> (\<forall>x. \<exists>y. R(x,y))"
+ \<longrightarrow> \<not> \<not> (\<forall>x. \<exists>y. R(x,y))\<close>
oops \<comment> \<open>NOT PROVED\<close>
text\<open>39\<close>
-lemma "\<not> (\<exists>x. \<forall>y. F(y,x) \<longleftrightarrow> \<not> F(y,y))"
+lemma \<open>\<not> (\<exists>x. \<forall>y. F(y,x) \<longleftrightarrow> \<not> F(y,y))\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>40. AMENDED\<close>
lemma
- "(\<exists>y. \<forall>x. F(x,y) \<longleftrightarrow> F(x,x)) \<longrightarrow>
- \<not> (\<forall>x. \<exists>y. \<forall>z. F(z,y) \<longleftrightarrow> \<not> F(z,x))"
+ \<open>(\<exists>y. \<forall>x. F(x,y) \<longleftrightarrow> F(x,x)) \<longrightarrow>
+ \<not> (\<forall>x. \<exists>y. \<forall>z. F(z,y) \<longleftrightarrow> \<not> F(z,x))\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>44\<close>
lemma
- "(\<forall>x. f(x) \<longrightarrow>
+ \<open>(\<forall>x. f(x) \<longrightarrow>
(\<exists>y. g(y) \<and> h(x,y) \<and> (\<exists>y. g(y) \<and> \<not> h(x,y)))) \<and>
(\<exists>x. j(x) \<and> (\<forall>y. g(y) \<longrightarrow> h(x,y)))
- \<longrightarrow> (\<exists>x. j(x) \<and> \<not> f(x))"
+ \<longrightarrow> (\<exists>x. j(x) \<and> \<not> f(x))\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>48\<close>
-lemma "(a = b \<or> c = d) \<and> (a = c \<or> b = d) \<longrightarrow> a = d \<or> b = c"
+lemma \<open>(a = b \<or> c = d) \<and> (a = c \<or> b = d) \<longrightarrow> a = d \<or> b = c\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>51\<close>
lemma
- "(\<exists>z w. \<forall>x y. P(x,y) \<longleftrightarrow> (x = z \<and> y = w)) \<longrightarrow>
- (\<exists>z. \<forall>x. \<exists>w. (\<forall>y. P(x,y) \<longleftrightarrow> y = w) \<longleftrightarrow> x = z)"
+ \<open>(\<exists>z w. \<forall>x y. P(x,y) \<longleftrightarrow> (x = z \<and> y = w)) \<longrightarrow>
+ (\<exists>z. \<forall>x. \<exists>w. (\<forall>y. P(x,y) \<longleftrightarrow> y = w) \<longleftrightarrow> x = z)\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>52\<close>
text \<open>Almost the same as 51.\<close>
lemma
- "(\<exists>z w. \<forall>x y. P(x,y) \<longleftrightarrow> (x = z \<and> y = w)) \<longrightarrow>
- (\<exists>w. \<forall>y. \<exists>z. (\<forall>x. P(x,y) \<longleftrightarrow> x = z) \<longleftrightarrow> y = w)"
+ \<open>(\<exists>z w. \<forall>x y. P(x,y) \<longleftrightarrow> (x = z \<and> y = w)) \<longrightarrow>
+ (\<exists>w. \<forall>y. \<exists>z. (\<forall>x. P(x,y) \<longleftrightarrow> x = z) \<longleftrightarrow> y = w)\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>56\<close>
-lemma "(\<forall>x. (\<exists>y. P(y) \<and> x = f(y)) \<longrightarrow> P(x)) \<longleftrightarrow> (\<forall>x. P(x) \<longrightarrow> P(f(x)))"
+lemma \<open>(\<forall>x. (\<exists>y. P(y) \<and> x = f(y)) \<longrightarrow> P(x)) \<longleftrightarrow> (\<forall>x. P(x) \<longrightarrow> P(f(x)))\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>57\<close>
lemma
- "P(f(a,b), f(b,c)) \<and> P(f(b,c), f(a,c)) \<and>
- (\<forall>x y z. P(x,y) \<and> P(y,z) \<longrightarrow> P(x,z)) \<longrightarrow> P(f(a,b), f(a,c))"
+ \<open>P(f(a,b), f(b,c)) \<and> P(f(b,c), f(a,c)) \<and>
+ (\<forall>x y z. P(x,y) \<and> P(y,z) \<longrightarrow> P(x,z)) \<longrightarrow> P(f(a,b), f(a,c))\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text\<open>60\<close>
-lemma "\<forall>x. P(x,f(x)) \<longleftrightarrow> (\<exists>y. (\<forall>z. P(z,y) \<longrightarrow> P(z,f(x))) \<and> P(x,y))"
+lemma \<open>\<forall>x. P(x,f(x)) \<longleftrightarrow> (\<exists>y. (\<forall>z. P(z,y) \<longrightarrow> P(z,f(x))) \<and> P(x,y))\<close>
by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
end
--- a/src/FOL/ex/Locale_Test/Locale_Test.thy Thu Jan 03 21:48:05 2019 +0100
+++ b/src/FOL/ex/Locale_Test/Locale_Test.thy Thu Jan 03 22:19:19 2019 +0100
@@ -16,9 +16,9 @@
lemmas less_mixin_thy_merge2 = le'.less_def
end
-lemma "gless(x, y) \<longleftrightarrow> gle(x, y) \<and> x \<noteq> y" (* mixin from first interpretation applied *)
+lemma \<open>gless(x, y) \<longleftrightarrow> gle(x, y) \<and> x \<noteq> y\<close> (* mixin from first interpretation applied *)
by (rule le1.less_mixin_thy_merge1)
-lemma "gless'(x, y) \<longleftrightarrow> gle'(x, y) \<and> x \<noteq> y" (* mixin from second interpretation applied *)
+lemma \<open>gless'(x, y) \<longleftrightarrow> gle'(x, y) \<and> x \<noteq> y\<close> (* mixin from second interpretation applied *)
by (rule le1.less_mixin_thy_merge2)
end
--- a/src/FOL/ex/Locale_Test/Locale_Test1.thy Thu Jan 03 21:48:05 2019 +0100
+++ b/src/FOL/ex/Locale_Test/Locale_Test1.thy Thu Jan 03 22:19:19 2019 +0100
@@ -9,24 +9,24 @@
begin
typedecl int
-instance int :: "term" ..
+instance int :: \<open>term\<close> ..
-consts plus :: "int => int => int" (infixl \<open>+\<close> 60)
- zero :: int (\<open>0\<close>)
- minus :: "int => int" (\<open>- _\<close>)
+consts plus :: \<open>int => int => int\<close> (infixl \<open>+\<close> 60)
+ zero :: \<open>int\<close> (\<open>0\<close>)
+ minus :: \<open>int => int\<close> (\<open>- _\<close>)
axiomatization where
- int_assoc: "(x + y::int) + z = x + (y + z)" and
- int_zero: "0 + x = x" and
- int_minus: "(-x) + x = 0" and
- int_minus2: "-(-x) = x"
+ int_assoc: \<open>(x + y::int) + z = x + (y + z)\<close> and
+ int_zero: \<open>0 + x = x\<close> and
+ int_minus: \<open>(-x) + x = 0\<close> and
+ int_minus2: \<open>-(-x) = x\<close>
section \<open>Inference of parameter types\<close>
locale param1 = fixes p
print_locale! param1
-locale param2 = fixes p :: 'b
+locale param2 = fixes p :: \<open>'b\<close>
print_locale! param2
(*
@@ -37,7 +37,7 @@
locale param3 = fixes p (infix \<open>..\<close> 50)
print_locale! param3
-locale param4 = fixes p :: "'a => 'a => 'a" (infix \<open>..\<close> 50)
+locale param4 = fixes p :: \<open>'a => 'a => 'a\<close> (infix \<open>..\<close> 50)
print_locale! param4
@@ -45,13 +45,13 @@
locale constraint1 =
fixes prod (infixl \<open>**\<close> 65)
- assumes l_id: "x ** y = x"
- assumes assoc: "(x ** y) ** z = x ** (y ** z)"
+ assumes l_id: \<open>x ** y = x\<close>
+ assumes assoc: \<open>(x ** y) ** z = x ** (y ** z)\<close>
print_locale! constraint1
locale constraint2 =
fixes p and q
- assumes "p = q"
+ assumes \<open>p = q\<close>
print_locale! constraint2
@@ -59,35 +59,35 @@
locale semi =
fixes prod (infixl \<open>**\<close> 65)
- assumes assoc: "(x ** y) ** z = x ** (y ** z)"
+ assumes assoc: \<open>(x ** y) ** z = x ** (y ** z)\<close>
print_locale! semi thm semi_def
locale lgrp = semi +
fixes one and inv
- assumes lone: "one ** x = x"
- and linv: "inv(x) ** x = one"
+ assumes lone: \<open>one ** x = x\<close>
+ and linv: \<open>inv(x) ** x = one\<close>
print_locale! lgrp thm lgrp_def lgrp_axioms_def
-locale add_lgrp = semi "(++)" for sum (infixl \<open>++\<close> 60) +
+locale add_lgrp = semi \<open>(++)\<close> for sum (infixl \<open>++\<close> 60) +
fixes zero and neg
- assumes lzero: "zero ++ x = x"
- and lneg: "neg(x) ++ x = zero"
+ assumes lzero: \<open>zero ++ x = x\<close>
+ and lneg: \<open>neg(x) ++ x = zero\<close>
print_locale! add_lgrp thm add_lgrp_def add_lgrp_axioms_def
-locale rev_lgrp = semi "%x y. y ++ x" for sum (infixl \<open>++\<close> 60)
+locale rev_lgrp = semi \<open>%x y. y ++ x\<close> for sum (infixl \<open>++\<close> 60)
print_locale! rev_lgrp thm rev_lgrp_def
-locale hom = f: semi f + g: semi g for f and g
+locale hom = f: semi \<open>f\<close> + g: semi \<open>g\<close> for f and g
print_locale! hom thm hom_def
-locale perturbation = semi + d: semi "%x y. delta(x) ** delta(y)" for delta
+locale perturbation = semi + d: semi \<open>%x y. delta(x) ** delta(y)\<close> for delta
print_locale! perturbation thm perturbation_def
-locale pert_hom = d1: perturbation f d1 + d2: perturbation f d2 for f d1 d2
+locale pert_hom = d1: perturbation \<open>f\<close> \<open>d1\<close> + d2: perturbation \<open>f\<close> \<open>d2\<close> for f d1 d2
print_locale! pert_hom thm pert_hom_def
text \<open>Alternative expression, obtaining nicer names in \<open>semi f\<close>.\<close>
-locale pert_hom' = semi f + d1: perturbation f d1 + d2: perturbation f d2 for f d1 d2
+locale pert_hom' = semi \<open>f\<close> + d1: perturbation \<open>f\<close> \<open>d1\<close> + d2: perturbation \<open>f\<close> \<open>d2\<close> for f d1 d2
print_locale! pert_hom' thm pert_hom'_def
@@ -96,32 +96,32 @@
locale logic =
fixes land (infixl \<open>&&\<close> 55)
and lnot (\<open>-- _\<close> [60] 60)
- assumes assoc: "(x && y) && z = x && (y && z)"
- and notnot: "-- (-- x) = x"
+ assumes assoc: \<open>(x && y) && z = x && (y && z)\<close>
+ and notnot: \<open>-- (-- x) = x\<close>
begin
definition lor (infixl \<open>||\<close> 50) where
- "x || y = --(-- x && -- y)"
+ \<open>x || y = --(-- x && -- y)\<close>
end
print_locale! logic
-locale use_decl = l: logic + semi "(||)"
+locale use_decl = l: logic + semi \<open>(||)\<close>
print_locale! use_decl thm use_decl_def
locale extra_type =
- fixes a :: 'a
- and P :: "'a => 'b => o"
+ fixes a :: \<open>'a\<close>
+ and P :: \<open>'a => 'b => o\<close>
begin
-definition test :: "'a => o"
- where "test(x) \<longleftrightarrow> (\<forall>b. P(x, b))"
+definition test :: \<open>'a => o\<close>
+ where \<open>test(x) \<longleftrightarrow> (\<forall>b. P(x, b))\<close>
end
-term extra_type.test thm extra_type.test_def
+term \<open>extra_type.test\<close> thm extra_type.test_def
-interpretation var?: extra_type "0" "%x y. x = 0" .
+interpretation var?: extra_type \<open>0\<close> \<open>%x y. x = 0\<close> .
thm var.test_def
@@ -143,12 +143,12 @@
declare [[show_hyps]]
locale "syntax" =
- fixes p1 :: "'a => 'b"
- and p2 :: "'b => o"
+ fixes p1 :: \<open>'a => 'b\<close>
+ and p2 :: \<open>'b => o\<close>
begin
-definition d1 :: "'a => o" (\<open>D1'(_')\<close>) where "d1(x) \<longleftrightarrow> \<not> p2(p1(x))"
-definition d2 :: "'b => o" (\<open>D2'(_')\<close>) where "d2(x) \<longleftrightarrow> \<not> p2(x)"
+definition d1 :: \<open>'a => o\<close> (\<open>D1'(_')\<close>) where \<open>d1(x) \<longleftrightarrow> \<not> p2(p1(x))\<close>
+definition d2 :: \<open>'b => o\<close> (\<open>D2'(_')\<close>) where \<open>d2(x) \<longleftrightarrow> \<not> p2(x)\<close>
thm d1_def d2_def
@@ -156,7 +156,7 @@
thm syntax.d1_def syntax.d2_def
-locale syntax' = "syntax" p1 p2 for p1 :: "'a => 'a" and p2 :: "'a => o"
+locale syntax' = "syntax" \<open>p1\<close> \<open>p2\<close> for p1 :: \<open>'a => 'a\<close> and p2 :: \<open>'a => o\<close>
begin
thm d1_def d2_def (* should print as "D1(?x) <-> ..." and "D2(?x) <-> ..." *)
@@ -168,7 +168,7 @@
end
-locale syntax'' = "syntax" p3 p2 for p3 :: "'a => 'b" and p2 :: "'b => o"
+locale syntax'' = "syntax" \<open>p3\<close> \<open>p2\<close> for p3 :: \<open>'a => 'b\<close> and p2 :: \<open>'b => o\<close>
begin
thm d1_def d2_def
@@ -194,14 +194,14 @@
fixes land (infixl \<open>&&\<close> 55)
and lor (infixl \<open>||\<close> 50)
and lnot (\<open>-- _\<close> [60] 60)
- assumes assoc: "(x && y) && z = x && (y && z)"
- and notnot: "-- (-- x) = x"
- defines "x || y == --(-- x && -- y)"
+ assumes assoc: \<open>(x && y) && z = x && (y && z)\<close>
+ and notnot: \<open>-- (-- x) = x\<close>
+ defines \<open>x || y == --(-- x && -- y)\<close>
begin
thm lor_def
-lemma "x || y = --(-- x && --y)"
+lemma \<open>x || y = --(-- x && --y)\<close>
by (unfold lor_def) (rule refl)
end
@@ -211,7 +211,7 @@
locale logic_def2 = logic_def
begin
-lemma "x || y = --(-- x && --y)"
+lemma \<open>x || y = --(-- x && --y)\<close>
by (unfold lor_def) (rule refl)
end
@@ -222,56 +222,56 @@
(* A somewhat arcane homomorphism example *)
definition semi_hom where
- "semi_hom(prod, sum, h) \<longleftrightarrow> (\<forall>x y. h(prod(x, y)) = sum(h(x), h(y)))"
+ \<open>semi_hom(prod, sum, h) \<longleftrightarrow> (\<forall>x y. h(prod(x, y)) = sum(h(x), h(y)))\<close>
lemma semi_hom_mult:
- "semi_hom(prod, sum, h) \<Longrightarrow> h(prod(x, y)) = sum(h(x), h(y))"
+ \<open>semi_hom(prod, sum, h) \<Longrightarrow> h(prod(x, y)) = sum(h(x), h(y))\<close>
by (simp add: semi_hom_def)
-locale semi_hom_loc = prod: semi prod + sum: semi sum
+locale semi_hom_loc = prod: semi \<open>prod\<close> + sum: semi \<open>sum\<close>
for prod and sum and h +
- assumes semi_homh: "semi_hom(prod, sum, h)"
+ assumes semi_homh: \<open>semi_hom(prod, sum, h)\<close>
notes semi_hom_mult = semi_hom_mult [OF semi_homh]
thm semi_hom_loc.semi_hom_mult
(* unspecified, attribute not applied in backgroud theory !!! *)
-lemma (in semi_hom_loc) "h(prod(x, y)) = sum(h(x), h(y))"
+lemma (in semi_hom_loc) \<open>h(prod(x, y)) = sum(h(x), h(y))\<close>
by (rule semi_hom_mult)
(* Referring to facts from within a context specification *)
lemma
- assumes x: "P \<longleftrightarrow> P"
+ assumes x: \<open>P \<longleftrightarrow> P\<close>
notes y = x
- shows True ..
+ shows \<open>True\<close> ..
section \<open>Theorem statements\<close>
lemma (in lgrp) lcancel:
- "x ** y = x ** z \<longleftrightarrow> y = z"
+ \<open>x ** y = x ** z \<longleftrightarrow> y = z\<close>
proof
- assume "x ** y = x ** z"
- then have "inv(x) ** x ** y = inv(x) ** x ** z" by (simp add: assoc)
- then show "y = z" by (simp add: lone linv)
+ assume \<open>x ** y = x ** z\<close>
+ then have \<open>inv(x) ** x ** y = inv(x) ** x ** z\<close> by (simp add: assoc)
+ then show \<open>y = z\<close> by (simp add: lone linv)
qed simp
print_locale! lgrp
locale rgrp = semi +
fixes one and inv
- assumes rone: "x ** one = x"
- and rinv: "x ** inv(x) = one"
+ assumes rone: \<open>x ** one = x\<close>
+ and rinv: \<open>x ** inv(x) = one\<close>
begin
lemma rcancel:
- "y ** x = z ** x \<longleftrightarrow> y = z"
+ \<open>y ** x = z ** x \<longleftrightarrow> y = z\<close>
proof
- assume "y ** x = z ** x"
- then have "y ** (x ** inv(x)) = z ** (x ** inv(x))"
+ assume \<open>y ** x = z ** x\<close>
+ then have \<open>y ** (x ** inv(x)) = z ** (x ** inv(x))\<close>
by (simp add: assoc [symmetric])
- then show "y = z" by (simp add: rone rinv)
+ then show \<open>y = z\<close> by (simp add: rone rinv)
qed simp
end
@@ -281,18 +281,18 @@
subsection \<open>Patterns\<close>
lemma (in rgrp)
- assumes "y ** x = z ** x" (is ?a)
- shows "y = z" (is ?t)
+ assumes \<open>y ** x = z ** x\<close> (is \<open>?a\<close>)
+ shows \<open>y = z\<close> (is \<open>?t\<close>)
proof -
txt \<open>Weird proof involving patterns from context element and conclusion.\<close>
{
- assume ?a
- then have "y ** (x ** inv(x)) = z ** (x ** inv(x))"
+ assume \<open>?a\<close>
+ then have \<open>y ** (x ** inv(x)) = z ** (x ** inv(x))\<close>
by (simp add: assoc [symmetric])
- then have ?t by (simp add: rone rinv)
+ then have \<open>?t\<close> by (simp add: rone rinv)
}
note x = this
- show ?t by (rule x [OF \<open>?a\<close>])
+ show \<open>?t\<close> by (rule x [OF \<open>?a\<close>])
qed
@@ -303,15 +303,15 @@
proof unfold_locales
{
fix x
- have "inv(x) ** x ** one = inv(x) ** x" by (simp add: linv lone)
- then show "x ** one = x" by (simp add: assoc lcancel)
+ have \<open>inv(x) ** x ** one = inv(x) ** x\<close> by (simp add: linv lone)
+ then show \<open>x ** one = x\<close> by (simp add: assoc lcancel)
}
note rone = this
{
fix x
- have "inv(x) ** x ** inv(x) = inv(x) ** one"
+ have \<open>inv(x) ** x ** inv(x) = inv(x) ** one\<close>
by (simp add: linv lone rone)
- then show "x ** inv(x) = one" by (simp add: assoc lcancel)
+ then show \<open>x ** inv(x) = one\<close> by (simp add: assoc lcancel)
}
qed
@@ -322,7 +322,7 @@
(* use of derived theorem *)
lemma (in lgrp)
- "y ** x = z ** x \<longleftrightarrow> y = z"
+ \<open>y ** x = z ** x \<longleftrightarrow> y = z\<close>
apply (rule rcancel)
done
@@ -332,15 +332,15 @@
proof unfold_locales
{
fix x
- have "one ** (x ** inv(x)) = x ** inv(x)" by (simp add: rinv rone)
- then show "one ** x = x" by (simp add: assoc [symmetric] rcancel)
+ have \<open>one ** (x ** inv(x)) = x ** inv(x)\<close> by (simp add: rinv rone)
+ then show \<open>one ** x = x\<close> by (simp add: assoc [symmetric] rcancel)
}
note lone = this
{
fix x
- have "inv(x) ** (x ** inv(x)) = one ** inv(x)"
+ have \<open>inv(x) ** (x ** inv(x)) = one ** inv(x)\<close>
by (simp add: rinv lone rone)
- then show "inv(x) ** x = one" by (simp add: assoc [symmetric] rcancel)
+ then show \<open>inv(x) ** x = one\<close> by (simp add: assoc [symmetric] rcancel)
}
qed
@@ -353,30 +353,30 @@
(* Duality *)
locale order =
- fixes less :: "'a => 'a => o" (infix \<open><<\<close> 50)
- assumes refl: "x << x"
- and trans: "[| x << y; y << z |] ==> x << z"
+ fixes less :: \<open>'a => 'a => o\<close> (infix \<open><<\<close> 50)
+ assumes refl: \<open>x << x\<close>
+ and trans: \<open>[| x << y; y << z |] ==> x << z\<close>
-sublocale order < dual: order "%x y. y << x"
+sublocale order < dual: order \<open>%x y. y << x\<close>
apply unfold_locales apply (rule refl) apply (blast intro: trans)
done
print_locale! order (* Only two instances of order. *)
locale order' =
- fixes less :: "'a => 'a => o" (infix \<open><<\<close> 50)
- assumes refl: "x << x"
- and trans: "[| x << y; y << z |] ==> x << z"
+ fixes less :: \<open>'a => 'a => o\<close> (infix \<open><<\<close> 50)
+ assumes refl: \<open>x << x\<close>
+ and trans: \<open>[| x << y; y << z |] ==> x << z\<close>
locale order_with_def = order'
begin
-definition greater :: "'a => 'a => o" (infix \<open>>>\<close> 50) where
- "x >> y \<longleftrightarrow> y << x"
+definition greater :: \<open>'a => 'a => o\<close> (infix \<open>>>\<close> 50) where
+ \<open>x >> y \<longleftrightarrow> y << x\<close>
end
-sublocale order_with_def < dual: order' "(>>)"
+sublocale order_with_def < dual: order' \<open>(>>)\<close>
apply unfold_locales
unfolding greater_def
apply (rule refl) apply (blast intro: trans)
@@ -392,17 +392,17 @@
locale A5 =
fixes A and B and C and D and E
- assumes eq: "A \<longleftrightarrow> B \<longleftrightarrow> C \<longleftrightarrow> D \<longleftrightarrow> E"
+ assumes eq: \<open>A \<longleftrightarrow> B \<longleftrightarrow> C \<longleftrightarrow> D \<longleftrightarrow> E\<close>
-sublocale A5 < 1: A5 _ _ D E C
+sublocale A5 < 1: A5 _ _ \<open>D\<close> \<open>E\<close> \<open>C\<close>
print_facts
using eq apply (blast intro: A5.intro) done
-sublocale A5 < 2: A5 C _ E _ A
+sublocale A5 < 2: A5 \<open>C\<close> _ \<open>E\<close> _ \<open>A\<close>
print_facts
using eq apply (blast intro: A5.intro) done
-sublocale A5 < 3: A5 B C A _ _
+sublocale A5 < 3: A5 \<open>B\<close> \<open>C\<close> \<open>A\<close> _ _
print_facts
using eq apply (blast intro: A5.intro) done
@@ -414,25 +414,25 @@
(* Free arguments of instance *)
locale trivial =
- fixes P and Q :: o
- assumes Q: "P \<longleftrightarrow> P \<longleftrightarrow> Q"
+ fixes P and Q :: \<open>o\<close>
+ assumes Q: \<open>P \<longleftrightarrow> P \<longleftrightarrow> Q\<close>
begin
-lemma Q_triv: "Q" using Q by fast
+lemma Q_triv: \<open>Q\<close> using Q by fast
end
-sublocale trivial < x: trivial x _
+sublocale trivial < x: trivial \<open>x\<close> _
apply unfold_locales using Q by fast
print_locale! trivial
context trivial
begin
-thm x.Q [where ?x = True]
+thm x.Q [where ?x = \<open>True\<close>]
end
-sublocale trivial < y: trivial Q Q
+sublocale trivial < y: trivial \<open>Q\<close> \<open>Q\<close>
by unfold_locales
(* Succeeds since previous interpretation is more general. *)
@@ -441,13 +441,13 @@
subsection \<open>Sublocale, then interpretation in theory\<close>
-interpretation int?: lgrp "(+)" "0" "minus"
+interpretation int?: lgrp \<open>(+)\<close> \<open>0\<close> \<open>minus\<close>
proof unfold_locales
qed (rule int_assoc int_zero int_minus)+
thm int.assoc int.semi_axioms
-interpretation int2?: semi "(+)"
+interpretation int2?: semi \<open>(+)\<close>
by unfold_locales (* subsumed, thm int2.assoc not generated *)
ML \<open>(Global_Theory.get_thms @{theory} "int2.assoc";
@@ -460,18 +460,18 @@
subsection \<open>Interpretation in theory, then sublocale\<close>
-interpretation fol: logic "(+)" "minus"
+interpretation fol: logic \<open>(+)\<close> \<open>minus\<close>
by unfold_locales (rule int_assoc int_minus2)+
locale logic2 =
fixes land (infixl \<open>&&\<close> 55)
and lnot (\<open>-- _\<close> [60] 60)
- assumes assoc: "(x && y) && z = x && (y && z)"
- and notnot: "-- (-- x) = x"
+ assumes assoc: \<open>(x && y) && z = x && (y && z)\<close>
+ and notnot: \<open>-- (-- x) = x\<close>
begin
definition lor (infixl \<open>||\<close> 50) where
- "x || y = --(-- x && -- y)"
+ \<open>x || y = --(-- x && -- y)\<close>
end
@@ -497,59 +497,59 @@
locale logic_o =
fixes land (infixl \<open>&&\<close> 55)
and lnot (\<open>-- _\<close> [60] 60)
- assumes assoc_o: "(x && y) && z \<longleftrightarrow> x && (y && z)"
- and notnot_o: "-- (-- x) \<longleftrightarrow> x"
+ assumes assoc_o: \<open>(x && y) && z \<longleftrightarrow> x && (y && z)\<close>
+ and notnot_o: \<open>-- (-- x) \<longleftrightarrow> x\<close>
begin
definition lor_o (infixl \<open>||\<close> 50) where
- "x || y \<longleftrightarrow> --(-- x && -- y)"
+ \<open>x || y \<longleftrightarrow> --(-- x && -- y)\<close>
end
-interpretation x: logic_o "(\<and>)" "Not"
- rewrites bool_logic_o: "x.lor_o(x, y) \<longleftrightarrow> x \<or> y"
+interpretation x: logic_o \<open>(\<and>)\<close> \<open>Not\<close>
+ rewrites bool_logic_o: \<open>x.lor_o(x, y) \<longleftrightarrow> x \<or> y\<close>
proof -
- show bool_logic_o: "PROP logic_o((\<and>), Not)" by unfold_locales fast+
- show "logic_o.lor_o((\<and>), Not, x, y) \<longleftrightarrow> x \<or> y"
+ show bool_logic_o: \<open>PROP logic_o((\<and>), Not)\<close> by unfold_locales fast+
+ show \<open>logic_o.lor_o((\<and>), Not, x, y) \<longleftrightarrow> x \<or> y\<close>
by (unfold logic_o.lor_o_def [OF bool_logic_o]) fast
qed
thm x.lor_o_def bool_logic_o
-lemma lor_triv: "z \<longleftrightarrow> z" ..
+lemma lor_triv: \<open>z \<longleftrightarrow> z\<close> ..
-lemma (in logic_o) lor_triv: "x || y \<longleftrightarrow> x || y" by fast
+lemma (in logic_o) lor_triv: \<open>x || y \<longleftrightarrow> x || y\<close> by fast
-thm lor_triv [where z = True] (* Check strict prefix. *)
+thm lor_triv [where z = \<open>True\<close>] (* Check strict prefix. *)
x.lor_triv
subsection \<open>Rewrite morphisms in expressions\<close>
-interpretation y: logic_o "(\<or>)" "Not" rewrites bool_logic_o2: "logic_o.lor_o((\<or>), Not, x, y) \<longleftrightarrow> x \<and> y"
+interpretation y: logic_o \<open>(\<or>)\<close> \<open>Not\<close> rewrites bool_logic_o2: \<open>logic_o.lor_o((\<or>), Not, x, y) \<longleftrightarrow> x \<and> y\<close>
proof -
- show bool_logic_o: "PROP logic_o((\<or>), Not)" by unfold_locales fast+
- show "logic_o.lor_o((\<or>), Not, x, y) \<longleftrightarrow> x \<and> y" unfolding logic_o.lor_o_def [OF bool_logic_o] by fast
+ show bool_logic_o: \<open>PROP logic_o((\<or>), Not)\<close> by unfold_locales fast+
+ show \<open>logic_o.lor_o((\<or>), Not, x, y) \<longleftrightarrow> x \<and> y\<close> unfolding logic_o.lor_o_def [OF bool_logic_o] by fast
qed
subsection \<open>Inheritance of rewrite morphisms\<close>
locale reflexive =
- fixes le :: "'a => 'a => o" (infix \<open>\<sqsubseteq>\<close> 50)
- assumes refl: "x \<sqsubseteq> x"
+ fixes le :: \<open>'a => 'a => o\<close> (infix \<open>\<sqsubseteq>\<close> 50)
+ assumes refl: \<open>x \<sqsubseteq> x\<close>
begin
-definition less (infix \<open>\<sqsubset>\<close> 50) where "x \<sqsubset> y \<longleftrightarrow> x \<sqsubseteq> y \<and> x \<noteq> y"
+definition less (infix \<open>\<sqsubset>\<close> 50) where \<open>x \<sqsubset> y \<longleftrightarrow> x \<sqsubseteq> y \<and> x \<noteq> y\<close>
end
axiomatization
- gle :: "'a => 'a => o" and gless :: "'a => 'a => o" and
- gle' :: "'a => 'a => o" and gless' :: "'a => 'a => o"
+ gle :: \<open>'a => 'a => o\<close> and gless :: \<open>'a => 'a => o\<close> and
+ gle' :: \<open>'a => 'a => o\<close> and gless' :: \<open>'a => 'a => o\<close>
where
- grefl: "gle(x, x)" and gless_def: "gless(x, y) \<longleftrightarrow> gle(x, y) \<and> x \<noteq> y" and
- grefl': "gle'(x, x)" and gless'_def: "gless'(x, y) \<longleftrightarrow> gle'(x, y) \<and> x \<noteq> y"
+ grefl: \<open>gle(x, x)\<close> and gless_def: \<open>gless(x, y) \<longleftrightarrow> gle(x, y) \<and> x \<noteq> y\<close> and
+ grefl': \<open>gle'(x, x)\<close> and gless'_def: \<open>gless'(x, y) \<longleftrightarrow> gle'(x, y) \<and> x \<noteq> y\<close>
text \<open>Setup\<close>
@@ -558,11 +558,11 @@
lemmas less_thm = less_def
end
-interpretation le: mixin gle rewrites "reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)"
+interpretation le: mixin \<open>gle\<close> rewrites \<open>reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)\<close>
proof -
- show "mixin(gle)" by unfold_locales (rule grefl)
+ show \<open>mixin(gle)\<close> by unfold_locales (rule grefl)
note reflexive = this[unfolded mixin_def]
- show "reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)"
+ show \<open>reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)\<close>
by (simp add: reflexive.less_def[OF reflexive] gless_def)
qed
@@ -573,11 +573,11 @@
lemmas less_thm2 = less_def
end
-interpretation le: mixin2 gle
+interpretation le: mixin2 \<open>gle\<close>
by unfold_locales
thm le.less_thm2 (* rewrite morphism applied *)
-lemma "gless(x, y) \<longleftrightarrow> gle(x, y) \<and> x \<noteq> y"
+lemma \<open>gless(x, y) \<longleftrightarrow> gle(x, y) \<and> x \<noteq> y\<close>
by (rule le.less_thm2)
text \<open>Rewrite morphism does not leak to a side branch.\<close>
@@ -587,11 +587,11 @@
lemmas less_thm3 = less_def
end
-interpretation le: mixin3 gle
+interpretation le: mixin3 \<open>gle\<close>
by unfold_locales
thm le.less_thm3 (* rewrite morphism not applied *)
-lemma "reflexive.less(gle, x, y) \<longleftrightarrow> gle(x, y) \<and> x \<noteq> y" by (rule le.less_thm3)
+lemma \<open>reflexive.less(gle, x, y) \<longleftrightarrow> gle(x, y) \<and> x \<noteq> y\<close> by (rule le.less_thm3)
text \<open>Rewrite morphism only available in original context\<close>
@@ -599,12 +599,12 @@
locale mixin4_mixin = mixin4_base
-interpretation le: mixin4_mixin gle
- rewrites "reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)"
+interpretation le: mixin4_mixin \<open>gle\<close>
+ rewrites \<open>reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)\<close>
proof -
- show "mixin4_mixin(gle)" by unfold_locales (rule grefl)
+ show \<open>mixin4_mixin(gle)\<close> by unfold_locales (rule grefl)
note reflexive = this[unfolded mixin4_mixin_def mixin4_base_def mixin_def]
- show "reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)"
+ show \<open>reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)\<close>
by (simp add: reflexive.less_def[OF reflexive] gless_def)
qed
@@ -613,16 +613,16 @@
lemmas less_thm4 = less_def
end
-locale mixin4_combined = le1?: mixin4_mixin le' + le2?: mixin4_copy le for le' le
+locale mixin4_combined = le1?: mixin4_mixin \<open>le'\<close> + le2?: mixin4_copy \<open>le\<close> for le' le
begin
lemmas less_thm4' = less_def
end
-interpretation le4: mixin4_combined gle' gle
+interpretation le4: mixin4_combined \<open>gle'\<close> \<open>gle\<close>
by unfold_locales (rule grefl')
thm le4.less_thm4' (* rewrite morphism not applied *)
-lemma "reflexive.less(gle, x, y) \<longleftrightarrow> gle(x, y) \<and> x \<noteq> y"
+lemma \<open>reflexive.less(gle, x, y) \<longleftrightarrow> gle(x, y) \<and> x \<noteq> y\<close>
by (rule le4.less_thm4')
text \<open>Inherited rewrite morphism applied to new theorem\<close>
@@ -631,22 +631,22 @@
locale mixin5_inherited = mixin5_base
-interpretation le5: mixin5_base gle
- rewrites "reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)"
+interpretation le5: mixin5_base \<open>gle\<close>
+ rewrites \<open>reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)\<close>
proof -
- show "mixin5_base(gle)" by unfold_locales
+ show \<open>mixin5_base(gle)\<close> by unfold_locales
note reflexive = this[unfolded mixin5_base_def mixin_def]
- show "reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)"
+ show \<open>reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)\<close>
by (simp add: reflexive.less_def[OF reflexive] gless_def)
qed
-interpretation le5: mixin5_inherited gle
+interpretation le5: mixin5_inherited \<open>gle\<close>
by unfold_locales
lemmas (in mixin5_inherited) less_thm5 = less_def
thm le5.less_thm5 (* rewrite morphism applied *)
-lemma "gless(x, y) \<longleftrightarrow> gle(x, y) \<and> x \<noteq> y"
+lemma \<open>gless(x, y) \<longleftrightarrow> gle(x, y) \<and> x \<noteq> y\<close>
by (rule le5.less_thm5)
text \<open>Rewrite morphism pushed down to existing inherited locale\<close>
@@ -655,23 +655,23 @@
locale mixin6_inherited = mixin5_base
-interpretation le6: mixin6_base gle
+interpretation le6: mixin6_base \<open>gle\<close>
by unfold_locales
-interpretation le6: mixin6_inherited gle
+interpretation le6: mixin6_inherited \<open>gle\<close>
by unfold_locales
-interpretation le6: mixin6_base gle
- rewrites "reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)"
+interpretation le6: mixin6_base \<open>gle\<close>
+ rewrites \<open>reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)\<close>
proof -
- show "mixin6_base(gle)" by unfold_locales
+ show \<open>mixin6_base(gle)\<close> by unfold_locales
note reflexive = this[unfolded mixin6_base_def mixin_def]
- show "reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)"
+ show \<open>reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)\<close>
by (simp add: reflexive.less_def[OF reflexive] gless_def)
qed
lemmas (in mixin6_inherited) less_thm6 = less_def
thm le6.less_thm6 (* mixin applied *)
-lemma "gless(x, y) \<longleftrightarrow> gle(x, y) \<and> x \<noteq> y"
+lemma \<open>gless(x, y) \<longleftrightarrow> gle(x, y) \<and> x \<noteq> y\<close>
by (rule le6.less_thm6)
text \<open>Existing rewrite morphism inherited through sublocale relation\<close>
@@ -680,22 +680,22 @@
locale mixin7_inherited = reflexive
-interpretation le7: mixin7_base gle
- rewrites "reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)"
+interpretation le7: mixin7_base \<open>gle\<close>
+ rewrites \<open>reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)\<close>
proof -
- show "mixin7_base(gle)" by unfold_locales
+ show \<open>mixin7_base(gle)\<close> by unfold_locales
note reflexive = this[unfolded mixin7_base_def mixin_def]
- show "reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)"
+ show \<open>reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)\<close>
by (simp add: reflexive.less_def[OF reflexive] gless_def)
qed
-interpretation le7: mixin7_inherited gle
+interpretation le7: mixin7_inherited \<open>gle\<close>
by unfold_locales
lemmas (in mixin7_inherited) less_thm7 = less_def
thm le7.less_thm7 (* before, rewrite morphism not applied *)
-lemma "reflexive.less(gle, x, y) \<longleftrightarrow> gle(x, y) \<and> x \<noteq> y"
+lemma \<open>reflexive.less(gle, x, y) \<longleftrightarrow> gle(x, y) \<and> x \<noteq> y\<close>
by (rule le7.less_thm7)
sublocale mixin7_inherited < mixin7_base
@@ -704,13 +704,13 @@
lemmas (in mixin7_inherited) less_thm7b = less_def
thm le7.less_thm7b (* after, mixin applied *)
-lemma "gless(x, y) \<longleftrightarrow> gle(x, y) \<and> x \<noteq> y"
+lemma \<open>gless(x, y) \<longleftrightarrow> gle(x, y) \<and> x \<noteq> y\<close>
by (rule le7.less_thm7b)
text \<open>This locale will be interpreted in later theories.\<close>
-locale mixin_thy_merge = le: reflexive le + le': reflexive le' for le le'
+locale mixin_thy_merge = le: reflexive \<open>le\<close> + le': reflexive \<open>le'\<close> for le le'
subsection \<open>Rewrite morphisms in sublocale\<close>
@@ -720,37 +720,37 @@
selection operator.\<close>
axiomatization glob_one and glob_inv
- where glob_lone: "prod(glob_one(prod), x) = x"
- and glob_linv: "prod(glob_inv(prod, x), x) = glob_one(prod)"
+ where glob_lone: \<open>prod(glob_one(prod), x) = x\<close>
+ and glob_linv: \<open>prod(glob_inv(prod, x), x) = glob_one(prod)\<close>
locale dgrp = semi
begin
-definition one where "one = glob_one(prod)"
+definition one where \<open>one = glob_one(prod)\<close>
-lemma lone: "one ** x = x"
+lemma lone: \<open>one ** x = x\<close>
unfolding one_def by (rule glob_lone)
-definition inv where "inv(x) = glob_inv(prod, x)"
+definition inv where \<open>inv(x) = glob_inv(prod, x)\<close>
-lemma linv: "inv(x) ** x = one"
+lemma linv: \<open>inv(x) ** x = one\<close>
unfolding one_def inv_def by (rule glob_linv)
end
sublocale lgrp < def?: dgrp
- rewrites one_equation: "dgrp.one(prod) = one" and inv_equation: "dgrp.inv(prod, x) = inv(x)"
+ rewrites one_equation: \<open>dgrp.one(prod) = one\<close> and inv_equation: \<open>dgrp.inv(prod, x) = inv(x)\<close>
proof -
- show "dgrp(prod)" by unfold_locales
+ show \<open>dgrp(prod)\<close> by unfold_locales
from this interpret d: dgrp .
\<comment> \<open>Unit\<close>
- have "dgrp.one(prod) = glob_one(prod)" by (rule d.one_def)
- also have "... = glob_one(prod) ** one" by (simp add: rone)
- also have "... = one" by (simp add: glob_lone)
- finally show "dgrp.one(prod) = one" .
+ have \<open>dgrp.one(prod) = glob_one(prod)\<close> by (rule d.one_def)
+ also have \<open>... = glob_one(prod) ** one\<close> by (simp add: rone)
+ also have \<open>... = one\<close> by (simp add: glob_lone)
+ finally show \<open>dgrp.one(prod) = one\<close> .
\<comment> \<open>Inverse\<close>
- then have "dgrp.inv(prod, x) ** x = inv(x) ** x" by (simp add: glob_linv d.linv linv)
- then show "dgrp.inv(prod, x) = inv(x)" by (simp add: rcancel)
+ then have \<open>dgrp.inv(prod, x) ** x = inv(x) ** x\<close> by (simp add: glob_linv d.linv linv)
+ then show \<open>dgrp.inv(prod, x) = inv(x)\<close> by (simp add: rcancel)
qed
print_locale! lgrp
@@ -760,35 +760,35 @@
text \<open>Equations stored in target\<close>
-lemma "dgrp.one(prod) = one" by (rule one_equation)
-lemma "dgrp.inv(prod, x) = inv(x)" by (rule inv_equation)
+lemma \<open>dgrp.one(prod) = one\<close> by (rule one_equation)
+lemma \<open>dgrp.inv(prod, x) = inv(x)\<close> by (rule inv_equation)
text \<open>Rewrite morphisms applied\<close>
-lemma "one = glob_one(prod)" by (rule one_def)
-lemma "inv(x) = glob_inv(prod, x)" by (rule inv_def)
+lemma \<open>one = glob_one(prod)\<close> by (rule one_def)
+lemma \<open>inv(x) = glob_inv(prod, x)\<close> by (rule inv_def)
end
text \<open>Interpreted versions\<close>
-lemma "0 = glob_one ((+))" by (rule int.def.one_def)
-lemma "- x = glob_inv((+), x)" by (rule int.def.inv_def)
+lemma \<open>0 = glob_one ((+))\<close> by (rule int.def.one_def)
+lemma \<open>- x = glob_inv((+), x)\<close> by (rule int.def.inv_def)
text \<open>Roundup applies rewrite morphisms at declaration level in DFS tree\<close>
-locale roundup = fixes x assumes true: "x \<longleftrightarrow> True"
+locale roundup = fixes x assumes true: \<open>x \<longleftrightarrow> True\<close>
-sublocale roundup \<subseteq> sub: roundup x rewrites "x \<longleftrightarrow> True \<and> True"
+sublocale roundup \<subseteq> sub: roundup \<open>x\<close> rewrites \<open>x \<longleftrightarrow> True \<and> True\<close>
apply unfold_locales apply (simp add: true) done
-lemma (in roundup) "True \<and> True \<longleftrightarrow> True" by (rule sub.true)
+lemma (in roundup) \<open>True \<and> True \<longleftrightarrow> True\<close> by (rule sub.true)
section \<open>Interpretation in named contexts\<close>
locale container
begin
-interpretation "private": roundup True by unfold_locales rule
+interpretation "private": roundup \<open>True\<close> by unfold_locales rule
lemmas true_copy = private.true
end
@@ -806,39 +806,39 @@
notepad
begin
- interpret "local": lgrp "(+)" "0" "minus"
+ interpret "local": lgrp \<open>(+)\<close> \<open>0\<close> \<open>minus\<close>
by unfold_locales (* subsumed *)
{
- fix zero :: int
- assume "!!x. zero + x = x" "!!x. (-x) + x = zero"
- then interpret local_fixed: lgrp "(+)" zero "minus"
+ fix zero :: \<open>int\<close>
+ assume \<open>!!x. zero + x = x\<close> \<open>!!x. (-x) + x = zero\<close>
+ then interpret local_fixed: lgrp \<open>(+)\<close> \<open>zero\<close> \<open>minus\<close>
by unfold_locales
thm local_fixed.lone
- print_dependencies! lgrp "(+)" 0 minus + lgrp "(+)" zero minus
+ print_dependencies! lgrp \<open>(+)\<close> \<open>0\<close> \<open>minus\<close> + lgrp \<open>(+)\<close> \<open>zero\<close> \<open>minus\<close>
}
- assume "!!x zero. zero + x = x" "!!x zero. (-x) + x = zero"
- then interpret local_free: lgrp "(+)" zero "minus" for zero
+ assume \<open>!!x zero. zero + x = x\<close> \<open>!!x zero. (-x) + x = zero\<close>
+ then interpret local_free: lgrp \<open>(+)\<close> \<open>zero\<close> \<open>minus\<close> for zero
by unfold_locales
- thm local_free.lone [where ?zero = 0]
+ thm local_free.lone [where ?zero = \<open>0\<close>]
end
notepad
begin
{
fix pand and pnot and por
- assume passoc: "\<And>x y z. pand(pand(x, y), z) \<longleftrightarrow> pand(x, pand(y, z))"
- and pnotnot: "\<And>x. pnot(pnot(x)) \<longleftrightarrow> x"
- and por_def: "\<And>x y. por(x, y) \<longleftrightarrow> pnot(pand(pnot(x), pnot(y)))"
- interpret loc: logic_o pand pnot
- rewrites por_eq: "\<And>x y. logic_o.lor_o(pand, pnot, x, y) \<longleftrightarrow> por(x, y)" (* FIXME *)
+ assume passoc: \<open>\<And>x y z. pand(pand(x, y), z) \<longleftrightarrow> pand(x, pand(y, z))\<close>
+ and pnotnot: \<open>\<And>x. pnot(pnot(x)) \<longleftrightarrow> x\<close>
+ and por_def: \<open>\<And>x y. por(x, y) \<longleftrightarrow> pnot(pand(pnot(x), pnot(y)))\<close>
+ interpret loc: logic_o \<open>pand\<close> \<open>pnot\<close>
+ rewrites por_eq: \<open>\<And>x y. logic_o.lor_o(pand, pnot, x, y) \<longleftrightarrow> por(x, y)\<close> (* FIXME *)
proof -
- show logic_o: "PROP logic_o(pand, pnot)" using passoc pnotnot by unfold_locales
+ show logic_o: \<open>PROP logic_o(pand, pnot)\<close> using passoc pnotnot by unfold_locales
fix x y
- show "logic_o.lor_o(pand, pnot, x, y) \<longleftrightarrow> por(x, y)"
+ show \<open>logic_o.lor_o(pand, pnot, x, y) \<longleftrightarrow> por(x, y)\<close>
by (unfold logic_o.lor_o_def [OF logic_o]) (rule por_def [symmetric])
qed
print_interps logic_o
- have "\<And>x y. por(x, y) \<longleftrightarrow> pnot(pand(pnot(x), pnot(y)))" by (rule loc.lor_o_def)
+ have \<open>\<And>x y. por(x, y) \<longleftrightarrow> pnot(pand(pnot(x), pnot(y)))\<close> by (rule loc.lor_o_def)
}
end
--- a/src/FOL/ex/Locale_Test/Locale_Test2.thy Thu Jan 03 21:48:05 2019 +0100
+++ b/src/FOL/ex/Locale_Test/Locale_Test2.thy Thu Jan 03 22:19:19 2019 +0100
@@ -8,12 +8,12 @@
imports Locale_Test1
begin
-interpretation le1: mixin_thy_merge gle gle'
- rewrites "reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)"
+interpretation le1: mixin_thy_merge \<open>gle\<close> \<open>gle'\<close>
+ rewrites \<open>reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)\<close>
proof -
- show "mixin_thy_merge(gle, gle')" by unfold_locales
+ show \<open>mixin_thy_merge(gle, gle')\<close> by unfold_locales
note reflexive = this[unfolded mixin_thy_merge_def, THEN conjunct1]
- show "reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)"
+ show \<open>reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)\<close>
by (simp add: reflexive.less_def[OF reflexive] gless_def)
qed
--- a/src/FOL/ex/Locale_Test/Locale_Test3.thy Thu Jan 03 21:48:05 2019 +0100
+++ b/src/FOL/ex/Locale_Test/Locale_Test3.thy Thu Jan 03 22:19:19 2019 +0100
@@ -8,12 +8,12 @@
imports Locale_Test1
begin
-interpretation le2: mixin_thy_merge gle gle'
- rewrites "reflexive.less(gle', x, y) \<longleftrightarrow> gless'(x, y)"
+interpretation le2: mixin_thy_merge \<open>gle\<close> \<open>gle'\<close>
+ rewrites \<open>reflexive.less(gle', x, y) \<longleftrightarrow> gless'(x, y)\<close>
proof -
- show "mixin_thy_merge(gle, gle')" by unfold_locales
+ show \<open>mixin_thy_merge(gle, gle')\<close> by unfold_locales
note reflexive = this[unfolded mixin_thy_merge_def, THEN conjunct2]
- show "reflexive.less(gle', x, y) \<longleftrightarrow> gless'(x, y)"
+ show \<open>reflexive.less(gle', x, y) \<longleftrightarrow> gless'(x, y)\<close>
by (simp add: reflexive.less_def[OF reflexive] gless'_def)
qed
--- a/src/FOL/ex/Miniscope.thy Thu Jan 03 21:48:05 2019 +0100
+++ b/src/FOL/ex/Miniscope.thy Thu Jan 03 22:19:19 2019 +0100
@@ -18,14 +18,14 @@
subsubsection \<open>de Morgan laws\<close>
lemma demorgans1:
- "\<not> (P \<and> Q) \<longleftrightarrow> \<not> P \<or> \<not> Q"
- "\<not> (P \<or> Q) \<longleftrightarrow> \<not> P \<and> \<not> Q"
- "\<not> \<not> P \<longleftrightarrow> P"
+ \<open>\<not> (P \<and> Q) \<longleftrightarrow> \<not> P \<or> \<not> Q\<close>
+ \<open>\<not> (P \<or> Q) \<longleftrightarrow> \<not> P \<and> \<not> Q\<close>
+ \<open>\<not> \<not> P \<longleftrightarrow> P\<close>
by blast+
lemma demorgans2:
- "\<And>P. \<not> (\<forall>x. P(x)) \<longleftrightarrow> (\<exists>x. \<not> P(x))"
- "\<And>P. \<not> (\<exists>x. P(x)) \<longleftrightarrow> (\<forall>x. \<not> P(x))"
+ \<open>\<And>P. \<not> (\<forall>x. P(x)) \<longleftrightarrow> (\<exists>x. \<not> P(x))\<close>
+ \<open>\<And>P. \<not> (\<exists>x. P(x)) \<longleftrightarrow> (\<forall>x. \<not> P(x))\<close>
by blast+
lemmas demorgans = demorgans1 demorgans2
@@ -33,10 +33,10 @@
(*** Removal of --> and <-> (positive and negative occurrences) ***)
(*Last one is important for computing a compact CNF*)
lemma nnf_simps:
- "(P \<longrightarrow> Q) \<longleftrightarrow> (\<not> P \<or> Q)"
- "\<not> (P \<longrightarrow> Q) \<longleftrightarrow> (P \<and> \<not> Q)"
- "(P \<longleftrightarrow> Q) \<longleftrightarrow> (\<not> P \<or> Q) \<and> (\<not> Q \<or> P)"
- "\<not> (P \<longleftrightarrow> Q) \<longleftrightarrow> (P \<or> Q) \<and> (\<not> P \<or> \<not> Q)"
+ \<open>(P \<longrightarrow> Q) \<longleftrightarrow> (\<not> P \<or> Q)\<close>
+ \<open>\<not> (P \<longrightarrow> Q) \<longleftrightarrow> (P \<and> \<not> Q)\<close>
+ \<open>(P \<longleftrightarrow> Q) \<longleftrightarrow> (\<not> P \<or> Q) \<and> (\<not> Q \<or> P)\<close>
+ \<open>\<not> (P \<longleftrightarrow> Q) \<longleftrightarrow> (P \<or> Q) \<and> (\<not> P \<or> \<not> Q)\<close>
by blast+
@@ -45,24 +45,24 @@
subsubsection \<open>Pushing in the existential quantifiers\<close>
lemma ex_simps:
- "(\<exists>x. P) \<longleftrightarrow> P"
- "\<And>P Q. (\<exists>x. P(x) \<and> Q) \<longleftrightarrow> (\<exists>x. P(x)) \<and> Q"
- "\<And>P Q. (\<exists>x. P \<and> Q(x)) \<longleftrightarrow> P \<and> (\<exists>x. Q(x))"
- "\<And>P Q. (\<exists>x. P(x) \<or> Q(x)) \<longleftrightarrow> (\<exists>x. P(x)) \<or> (\<exists>x. Q(x))"
- "\<And>P Q. (\<exists>x. P(x) \<or> Q) \<longleftrightarrow> (\<exists>x. P(x)) \<or> Q"
- "\<And>P Q. (\<exists>x. P \<or> Q(x)) \<longleftrightarrow> P \<or> (\<exists>x. Q(x))"
+ \<open>(\<exists>x. P) \<longleftrightarrow> P\<close>
+ \<open>\<And>P Q. (\<exists>x. P(x) \<and> Q) \<longleftrightarrow> (\<exists>x. P(x)) \<and> Q\<close>
+ \<open>\<And>P Q. (\<exists>x. P \<and> Q(x)) \<longleftrightarrow> P \<and> (\<exists>x. Q(x))\<close>
+ \<open>\<And>P Q. (\<exists>x. P(x) \<or> Q(x)) \<longleftrightarrow> (\<exists>x. P(x)) \<or> (\<exists>x. Q(x))\<close>
+ \<open>\<And>P Q. (\<exists>x. P(x) \<or> Q) \<longleftrightarrow> (\<exists>x. P(x)) \<or> Q\<close>
+ \<open>\<And>P Q. (\<exists>x. P \<or> Q(x)) \<longleftrightarrow> P \<or> (\<exists>x. Q(x))\<close>
by blast+
subsubsection \<open>Pushing in the universal quantifiers\<close>
lemma all_simps:
- "(\<forall>x. P) \<longleftrightarrow> P"
- "\<And>P Q. (\<forall>x. P(x) \<and> Q(x)) \<longleftrightarrow> (\<forall>x. P(x)) \<and> (\<forall>x. Q(x))"
- "\<And>P Q. (\<forall>x. P(x) \<and> Q) \<longleftrightarrow> (\<forall>x. P(x)) \<and> Q"
- "\<And>P Q. (\<forall>x. P \<and> Q(x)) \<longleftrightarrow> P \<and> (\<forall>x. Q(x))"
- "\<And>P Q. (\<forall>x. P(x) \<or> Q) \<longleftrightarrow> (\<forall>x. P(x)) \<or> Q"
- "\<And>P Q. (\<forall>x. P \<or> Q(x)) \<longleftrightarrow> P \<or> (\<forall>x. Q(x))"
+ \<open>(\<forall>x. P) \<longleftrightarrow> P\<close>
+ \<open>\<And>P Q. (\<forall>x. P(x) \<and> Q(x)) \<longleftrightarrow> (\<forall>x. P(x)) \<and> (\<forall>x. Q(x))\<close>
+ \<open>\<And>P Q. (\<forall>x. P(x) \<and> Q) \<longleftrightarrow> (\<forall>x. P(x)) \<and> Q\<close>
+ \<open>\<And>P Q. (\<forall>x. P \<and> Q(x)) \<longleftrightarrow> P \<and> (\<forall>x. Q(x))\<close>
+ \<open>\<And>P Q. (\<forall>x. P(x) \<or> Q) \<longleftrightarrow> (\<forall>x. P(x)) \<or> Q\<close>
+ \<open>\<And>P Q. (\<forall>x. P \<or> Q(x)) \<longleftrightarrow> P \<or> (\<forall>x. Q(x))\<close>
by blast+
lemmas mini_simps = demorgans nnf_simps ex_simps all_simps
--- a/src/FOL/ex/Nat.thy Thu Jan 03 21:48:05 2019 +0100
+++ b/src/FOL/ex/Nat.thy Thu Jan 03 22:19:19 2019 +0100
@@ -10,27 +10,27 @@
begin
typedecl nat
-instance nat :: "term" ..
+instance nat :: \<open>term\<close> ..
axiomatization
- Zero :: nat (\<open>0\<close>) and
- Suc :: "nat => nat" and
- rec :: "[nat, 'a, [nat, 'a] => 'a] => 'a"
+ Zero :: \<open>nat\<close> (\<open>0\<close>) and
+ Suc :: \<open>nat => nat\<close> and
+ rec :: \<open>[nat, 'a, [nat, 'a] => 'a] => 'a\<close>
where
- induct: "[| P(0); !!x. P(x) ==> P(Suc(x)) |] ==> P(n)" and
- Suc_inject: "Suc(m)=Suc(n) ==> m=n" and
- Suc_neq_0: "Suc(m)=0 ==> R" and
- rec_0: "rec(0,a,f) = a" and
- rec_Suc: "rec(Suc(m), a, f) = f(m, rec(m,a,f))"
+ induct: \<open>[| P(0); !!x. P(x) ==> P(Suc(x)) |] ==> P(n)\<close> and
+ Suc_inject: \<open>Suc(m)=Suc(n) ==> m=n\<close> and
+ Suc_neq_0: \<open>Suc(m)=0 ==> R\<close> and
+ rec_0: \<open>rec(0,a,f) = a\<close> and
+ rec_Suc: \<open>rec(Suc(m), a, f) = f(m, rec(m,a,f))\<close>
-definition add :: "[nat, nat] => nat" (infixl \<open>+\<close> 60)
- where "m + n == rec(m, n, %x y. Suc(y))"
+definition add :: \<open>[nat, nat] => nat\<close> (infixl \<open>+\<close> 60)
+ where \<open>m + n == rec(m, n, %x y. Suc(y))\<close>
subsection \<open>Proofs about the natural numbers\<close>
-lemma Suc_n_not_n: "Suc(k) \<noteq> k"
-apply (rule_tac n = k in induct)
+lemma Suc_n_not_n: \<open>Suc(k) \<noteq> k\<close>
+apply (rule_tac n = \<open>k\<close> in induct)
apply (rule notI)
apply (erule Suc_neq_0)
apply (rule notI)
@@ -38,7 +38,7 @@
apply (erule Suc_inject)
done
-lemma "(k+m)+n = k+(m+n)"
+lemma \<open>(k+m)+n = k+(m+n)\<close>
apply (rule induct)
back
back
@@ -48,37 +48,37 @@
back
oops
-lemma add_0 [simp]: "0+n = n"
+lemma add_0 [simp]: \<open>0+n = n\<close>
apply (unfold add_def)
apply (rule rec_0)
done
-lemma add_Suc [simp]: "Suc(m)+n = Suc(m+n)"
+lemma add_Suc [simp]: \<open>Suc(m)+n = Suc(m+n)\<close>
apply (unfold add_def)
apply (rule rec_Suc)
done
-lemma add_assoc: "(k+m)+n = k+(m+n)"
-apply (rule_tac n = k in induct)
+lemma add_assoc: \<open>(k+m)+n = k+(m+n)\<close>
+apply (rule_tac n = \<open>k\<close> in induct)
apply simp
apply simp
done
-lemma add_0_right: "m+0 = m"
-apply (rule_tac n = m in induct)
+lemma add_0_right: \<open>m+0 = m\<close>
+apply (rule_tac n = \<open>m\<close> in induct)
apply simp
apply simp
done
-lemma add_Suc_right: "m+Suc(n) = Suc(m+n)"
-apply (rule_tac n = m in induct)
+lemma add_Suc_right: \<open>m+Suc(n) = Suc(m+n)\<close>
+apply (rule_tac n = \<open>m\<close> in induct)
apply simp_all
done
lemma
- assumes prem: "!!n. f(Suc(n)) = Suc(f(n))"
- shows "f(i+j) = i+f(j)"
-apply (rule_tac n = i in induct)
+ assumes prem: \<open>!!n. f(Suc(n)) = Suc(f(n))\<close>
+ shows \<open>f(i+j) = i+f(j)\<close>
+apply (rule_tac n = \<open>i\<close> in induct)
apply simp
apply (simp add: prem)
done
--- a/src/FOL/ex/Nat_Class.thy Thu Jan 03 21:48:05 2019 +0100
+++ b/src/FOL/ex/Nat_Class.thy Thu Jan 03 22:19:19 2019 +0100
@@ -16,21 +16,21 @@
\<close>
class nat =
- fixes Zero :: 'a (\<open>0\<close>)
- and Suc :: "'a \<Rightarrow> 'a"
- and rec :: "'a \<Rightarrow> 'a \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'a"
- assumes induct: "P(0) \<Longrightarrow> (\<And>x. P(x) \<Longrightarrow> P(Suc(x))) \<Longrightarrow> P(n)"
- and Suc_inject: "Suc(m) = Suc(n) \<Longrightarrow> m = n"
- and Suc_neq_Zero: "Suc(m) = 0 \<Longrightarrow> R"
- and rec_Zero: "rec(0, a, f) = a"
- and rec_Suc: "rec(Suc(m), a, f) = f(m, rec(m, a, f))"
+ fixes Zero :: \<open>'a\<close> (\<open>0\<close>)
+ and Suc :: \<open>'a \<Rightarrow> 'a\<close>
+ and rec :: \<open>'a \<Rightarrow> 'a \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'a\<close>
+ assumes induct: \<open>P(0) \<Longrightarrow> (\<And>x. P(x) \<Longrightarrow> P(Suc(x))) \<Longrightarrow> P(n)\<close>
+ and Suc_inject: \<open>Suc(m) = Suc(n) \<Longrightarrow> m = n\<close>
+ and Suc_neq_Zero: \<open>Suc(m) = 0 \<Longrightarrow> R\<close>
+ and rec_Zero: \<open>rec(0, a, f) = a\<close>
+ and rec_Suc: \<open>rec(Suc(m), a, f) = f(m, rec(m, a, f))\<close>
begin
-definition add :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl \<open>+\<close> 60)
- where "m + n = rec(m, n, \<lambda>x y. Suc(y))"
+definition add :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close> (infixl \<open>+\<close> 60)
+ where \<open>m + n = rec(m, n, \<lambda>x y. Suc(y))\<close>
-lemma Suc_n_not_n: "Suc(k) \<noteq> (k::'a)"
- apply (rule_tac n = k in induct)
+lemma Suc_n_not_n: \<open>Suc(k) \<noteq> (k::'a)\<close>
+ apply (rule_tac n = \<open>k\<close> in induct)
apply (rule notI)
apply (erule Suc_neq_Zero)
apply (rule notI)
@@ -38,7 +38,7 @@
apply (erule Suc_inject)
done
-lemma "(k + m) + n = k + (m + n)"
+lemma \<open>(k + m) + n = k + (m + n)\<close>
apply (rule induct)
back
back
@@ -47,37 +47,37 @@
back
oops
-lemma add_Zero [simp]: "0 + n = n"
+lemma add_Zero [simp]: \<open>0 + n = n\<close>
apply (unfold add_def)
apply (rule rec_Zero)
done
-lemma add_Suc [simp]: "Suc(m) + n = Suc(m + n)"
+lemma add_Suc [simp]: \<open>Suc(m) + n = Suc(m + n)\<close>
apply (unfold add_def)
apply (rule rec_Suc)
done
-lemma add_assoc: "(k + m) + n = k + (m + n)"
- apply (rule_tac n = k in induct)
+lemma add_assoc: \<open>(k + m) + n = k + (m + n)\<close>
+ apply (rule_tac n = \<open>k\<close> in induct)
apply simp
apply simp
done
-lemma add_Zero_right: "m + 0 = m"
- apply (rule_tac n = m in induct)
+lemma add_Zero_right: \<open>m + 0 = m\<close>
+ apply (rule_tac n = \<open>m\<close> in induct)
apply simp
apply simp
done
-lemma add_Suc_right: "m + Suc(n) = Suc(m + n)"
- apply (rule_tac n = m in induct)
+lemma add_Suc_right: \<open>m + Suc(n) = Suc(m + n)\<close>
+ apply (rule_tac n = \<open>m\<close> in induct)
apply simp_all
done
lemma
- assumes prem: "\<And>n. f(Suc(n)) = Suc(f(n))"
- shows "f(i + j) = i + f(j)"
- apply (rule_tac n = i in induct)
+ assumes prem: \<open>\<And>n. f(Suc(n)) = Suc(f(n))\<close>
+ shows \<open>f(i + j) = i + f(j)\<close>
+ apply (rule_tac n = \<open>i\<close> in induct)
apply simp
apply (simp add: prem)
done
--- a/src/FOL/ex/Natural_Numbers.thy Thu Jan 03 21:48:05 2019 +0100
+++ b/src/FOL/ex/Natural_Numbers.thy Thu Jan 03 22:19:19 2019 +0100
@@ -14,59 +14,59 @@
\<close>
typedecl nat
-instance nat :: "term" ..
+instance nat :: \<open>term\<close> ..
axiomatization
- Zero :: nat (\<open>0\<close>) and
- Suc :: "nat => nat" and
- rec :: "[nat, 'a, [nat, 'a] => 'a] => 'a"
+ Zero :: \<open>nat\<close> (\<open>0\<close>) and
+ Suc :: \<open>nat => nat\<close> and
+ rec :: \<open>[nat, 'a, [nat, 'a] => 'a] => 'a\<close>
where
induct [case_names 0 Suc, induct type: nat]:
- "P(0) ==> (!!x. P(x) ==> P(Suc(x))) ==> P(n)" and
- Suc_inject: "Suc(m) = Suc(n) ==> m = n" and
- Suc_neq_0: "Suc(m) = 0 ==> R" and
- rec_0: "rec(0, a, f) = a" and
- rec_Suc: "rec(Suc(m), a, f) = f(m, rec(m, a, f))"
+ \<open>P(0) ==> (!!x. P(x) ==> P(Suc(x))) ==> P(n)\<close> and
+ Suc_inject: \<open>Suc(m) = Suc(n) ==> m = n\<close> and
+ Suc_neq_0: \<open>Suc(m) = 0 ==> R\<close> and
+ rec_0: \<open>rec(0, a, f) = a\<close> and
+ rec_Suc: \<open>rec(Suc(m), a, f) = f(m, rec(m, a, f))\<close>
-lemma Suc_n_not_n: "Suc(k) \<noteq> k"
-proof (induct k)
- show "Suc(0) \<noteq> 0"
+lemma Suc_n_not_n: \<open>Suc(k) \<noteq> k\<close>
+proof (induct \<open>k\<close>)
+ show \<open>Suc(0) \<noteq> 0\<close>
proof
- assume "Suc(0) = 0"
- then show False by (rule Suc_neq_0)
+ assume \<open>Suc(0) = 0\<close>
+ then show \<open>False\<close> by (rule Suc_neq_0)
qed
next
- fix n assume hyp: "Suc(n) \<noteq> n"
- show "Suc(Suc(n)) \<noteq> Suc(n)"
+ fix n assume hyp: \<open>Suc(n) \<noteq> n\<close>
+ show \<open>Suc(Suc(n)) \<noteq> Suc(n)\<close>
proof
- assume "Suc(Suc(n)) = Suc(n)"
- then have "Suc(n) = n" by (rule Suc_inject)
- with hyp show False by contradiction
+ assume \<open>Suc(Suc(n)) = Suc(n)\<close>
+ then have \<open>Suc(n) = n\<close> by (rule Suc_inject)
+ with hyp show \<open>False\<close> by contradiction
qed
qed
-definition add :: "nat => nat => nat" (infixl \<open>+\<close> 60)
- where "m + n = rec(m, n, \<lambda>x y. Suc(y))"
+definition add :: \<open>nat => nat => nat\<close> (infixl \<open>+\<close> 60)
+ where \<open>m + n = rec(m, n, \<lambda>x y. Suc(y))\<close>
-lemma add_0 [simp]: "0 + n = n"
+lemma add_0 [simp]: \<open>0 + n = n\<close>
unfolding add_def by (rule rec_0)
-lemma add_Suc [simp]: "Suc(m) + n = Suc(m + n)"
+lemma add_Suc [simp]: \<open>Suc(m) + n = Suc(m + n)\<close>
unfolding add_def by (rule rec_Suc)
-lemma add_assoc: "(k + m) + n = k + (m + n)"
- by (induct k) simp_all
+lemma add_assoc: \<open>(k + m) + n = k + (m + n)\<close>
+ by (induct \<open>k\<close>) simp_all
-lemma add_0_right: "m + 0 = m"
- by (induct m) simp_all
+lemma add_0_right: \<open>m + 0 = m\<close>
+ by (induct \<open>m\<close>) simp_all
-lemma add_Suc_right: "m + Suc(n) = Suc(m + n)"
- by (induct m) simp_all
+lemma add_Suc_right: \<open>m + Suc(n) = Suc(m + n)\<close>
+ by (induct \<open>m\<close>) simp_all
lemma
- assumes "!!n. f(Suc(n)) = Suc(f(n))"
- shows "f(i + j) = i + f(j)"
- using assms by (induct i) simp_all
+ assumes \<open>!!n. f(Suc(n)) = Suc(f(n))\<close>
+ shows \<open>f(i + j) = i + f(j)\<close>
+ using assms by (induct \<open>i\<close>) simp_all
end
--- a/src/FOL/ex/Prolog.thy Thu Jan 03 21:48:05 2019 +0100
+++ b/src/FOL/ex/Prolog.thy Thu Jan 03 22:19:19 2019 +0100
@@ -10,31 +10,31 @@
begin
typedecl 'a list
-instance list :: ("term") "term" ..
+instance list :: (\<open>term\<close>) \<open>term\<close> ..
axiomatization
- Nil :: "'a list" and
- Cons :: "['a, 'a list]=> 'a list" (infixr \<open>:\<close> 60) and
- app :: "['a list, 'a list, 'a list] => o" and
- rev :: "['a list, 'a list] => o"
+ Nil :: \<open>'a list\<close> and
+ Cons :: \<open>['a, 'a list]=> 'a list\<close> (infixr \<open>:\<close> 60) and
+ app :: \<open>['a list, 'a list, 'a list] => o\<close> and
+ rev :: \<open>['a list, 'a list] => o\<close>
where
- appNil: "app(Nil,ys,ys)" and
- appCons: "app(xs,ys,zs) ==> app(x:xs, ys, x:zs)" and
- revNil: "rev(Nil,Nil)" and
- revCons: "[| rev(xs,ys); app(ys, x:Nil, zs) |] ==> rev(x:xs, zs)"
+ appNil: \<open>app(Nil,ys,ys)\<close> and
+ appCons: \<open>app(xs,ys,zs) ==> app(x:xs, ys, x:zs)\<close> and
+ revNil: \<open>rev(Nil,Nil)\<close> and
+ revCons: \<open>[| rev(xs,ys); app(ys, x:Nil, zs) |] ==> rev(x:xs, zs)\<close>
-schematic_goal "app(a:b:c:Nil, d:e:Nil, ?x)"
+schematic_goal \<open>app(a:b:c:Nil, d:e:Nil, ?x)\<close>
apply (rule appNil appCons)
apply (rule appNil appCons)
apply (rule appNil appCons)
apply (rule appNil appCons)
done
-schematic_goal "app(?x, c:d:Nil, a:b:c:d:Nil)"
+schematic_goal \<open>app(?x, c:d:Nil, a:b:c:d:Nil)\<close>
apply (rule appNil appCons)+
done
-schematic_goal "app(?x, ?y, a:b:c:d:Nil)"
+schematic_goal \<open>app(?x, ?y, a:b:c:d:Nil)\<close>
apply (rule appNil appCons)+
back
back
@@ -47,15 +47,15 @@
lemmas rules = appNil appCons revNil revCons
-schematic_goal "rev(a:b:c:d:Nil, ?x)"
+schematic_goal \<open>rev(a:b:c:d:Nil, ?x)\<close>
apply (rule rules)+
done
-schematic_goal "rev(a:b:c:d:e:f:g:h:i:j:k:l:m:n:Nil, ?w)"
+schematic_goal \<open>rev(a:b:c:d:e:f:g:h:i:j:k:l:m:n:Nil, ?w)\<close>
apply (rule rules)+
done
-schematic_goal "rev(?x, a:b:c:Nil)"
+schematic_goal \<open>rev(?x, a:b:c:Nil)\<close>
apply (rule rules)+ \<comment> \<open>does not solve it directly!\<close>
back
back
@@ -67,23 +67,23 @@
DEPTH_FIRST (has_fewer_prems 1) (resolve_tac ctxt @{thms rules} 1)
\<close>
-schematic_goal "rev(?x, a:b:c:Nil)"
+schematic_goal \<open>rev(?x, a:b:c:Nil)\<close>
apply (tactic \<open>prolog_tac @{context}\<close>)
done
-schematic_goal "rev(a:?x:c:?y:Nil, d:?z:b:?u)"
+schematic_goal \<open>rev(a:?x:c:?y:Nil, d:?z:b:?u)\<close>
apply (tactic \<open>prolog_tac @{context}\<close>)
done
(*rev([a..p], ?w) requires 153 inferences *)
-schematic_goal "rev(a:b:c:d:e:f:g:h:i:j:k:l:m:n:o:p:Nil, ?w)"
+schematic_goal \<open>rev(a:b:c:d:e:f:g:h:i:j:k:l:m:n:o:p:Nil, ?w)\<close>
apply (tactic \<open>
DEPTH_SOLVE (resolve_tac @{context} ([@{thm refl}, @{thm conjI}] @ @{thms rules}) 1)\<close>)
done
(*?x has 16, ?y has 32; rev(?y,?w) requires 561 (rather large) inferences
total inferences = 2 + 1 + 17 + 561 = 581*)
-schematic_goal "a:b:c:d:e:f:g:h:i:j:k:l:m:n:o:p:Nil = ?x \<and> app(?x,?x,?y) \<and> rev(?y,?w)"
+schematic_goal \<open>a:b:c:d:e:f:g:h:i:j:k:l:m:n:o:p:Nil = ?x \<and> app(?x,?x,?y) \<and> rev(?y,?w)\<close>
apply (tactic \<open>
DEPTH_SOLVE (resolve_tac @{context} ([@{thm refl}, @{thm conjI}] @ @{thms rules}) 1)\<close>)
done
--- a/src/FOL/ex/Propositional_Cla.thy Thu Jan 03 21:48:05 2019 +0100
+++ b/src/FOL/ex/Propositional_Cla.thy Thu Jan 03 22:19:19 2019 +0100
@@ -11,110 +11,110 @@
text \<open>commutative laws of \<open>\<and>\<close> and \<open>\<or>\<close>\<close>
-lemma "P \<and> Q \<longrightarrow> Q \<and> P"
+lemma \<open>P \<and> Q \<longrightarrow> Q \<and> P\<close>
by (tactic "IntPr.fast_tac @{context} 1")
-lemma "P \<or> Q \<longrightarrow> Q \<or> P"
+lemma \<open>P \<or> Q \<longrightarrow> Q \<or> P\<close>
by fast
text \<open>associative laws of \<open>\<and>\<close> and \<open>\<or>\<close>\<close>
-lemma "(P \<and> Q) \<and> R \<longrightarrow> P \<and> (Q \<and> R)"
+lemma \<open>(P \<and> Q) \<and> R \<longrightarrow> P \<and> (Q \<and> R)\<close>
by fast
-lemma "(P \<or> Q) \<or> R \<longrightarrow> P \<or> (Q \<or> R)"
+lemma \<open>(P \<or> Q) \<or> R \<longrightarrow> P \<or> (Q \<or> R)\<close>
by fast
text \<open>distributive laws of \<open>\<and>\<close> and \<open>\<or>\<close>\<close>
-lemma "(P \<and> Q) \<or> R \<longrightarrow> (P \<or> R) \<and> (Q \<or> R)"
+lemma \<open>(P \<and> Q) \<or> R \<longrightarrow> (P \<or> R) \<and> (Q \<or> R)\<close>
by fast
-lemma "(P \<or> R) \<and> (Q \<or> R) \<longrightarrow> (P \<and> Q) \<or> R"
+lemma \<open>(P \<or> R) \<and> (Q \<or> R) \<longrightarrow> (P \<and> Q) \<or> R\<close>
by fast
-lemma "(P \<or> Q) \<and> R \<longrightarrow> (P \<and> R) \<or> (Q \<and> R)"
+lemma \<open>(P \<or> Q) \<and> R \<longrightarrow> (P \<and> R) \<or> (Q \<and> R)\<close>
by fast
-lemma "(P \<and> R) \<or> (Q \<and> R) \<longrightarrow> (P \<or> Q) \<and> R"
+lemma \<open>(P \<and> R) \<or> (Q \<and> R) \<longrightarrow> (P \<or> Q) \<and> R\<close>
by fast
text \<open>Laws involving implication\<close>
-lemma "(P \<longrightarrow> R) \<and> (Q \<longrightarrow> R) \<longleftrightarrow> (P \<or> Q \<longrightarrow> R)"
+lemma \<open>(P \<longrightarrow> R) \<and> (Q \<longrightarrow> R) \<longleftrightarrow> (P \<or> Q \<longrightarrow> R)\<close>
by fast
-lemma "(P \<and> Q \<longrightarrow> R) \<longleftrightarrow> (P \<longrightarrow> (Q \<longrightarrow> R))"
+lemma \<open>(P \<and> Q \<longrightarrow> R) \<longleftrightarrow> (P \<longrightarrow> (Q \<longrightarrow> R))\<close>
by fast
-lemma "((P \<longrightarrow> R) \<longrightarrow> R) \<longrightarrow> ((Q \<longrightarrow> R) \<longrightarrow> R) \<longrightarrow> (P \<and> Q \<longrightarrow> R) \<longrightarrow> R"
+lemma \<open>((P \<longrightarrow> R) \<longrightarrow> R) \<longrightarrow> ((Q \<longrightarrow> R) \<longrightarrow> R) \<longrightarrow> (P \<and> Q \<longrightarrow> R) \<longrightarrow> R\<close>
by fast
-lemma "\<not> (P \<longrightarrow> R) \<longrightarrow> \<not> (Q \<longrightarrow> R) \<longrightarrow> \<not> (P \<and> Q \<longrightarrow> R)"
+lemma \<open>\<not> (P \<longrightarrow> R) \<longrightarrow> \<not> (Q \<longrightarrow> R) \<longrightarrow> \<not> (P \<and> Q \<longrightarrow> R)\<close>
by fast
-lemma "(P \<longrightarrow> Q \<and> R) \<longleftrightarrow> (P \<longrightarrow> Q) \<and> (P \<longrightarrow> R)"
+lemma \<open>(P \<longrightarrow> Q \<and> R) \<longleftrightarrow> (P \<longrightarrow> Q) \<and> (P \<longrightarrow> R)\<close>
by fast
text \<open>Propositions-as-types\<close>
\<comment> \<open>The combinator K\<close>
-lemma "P \<longrightarrow> (Q \<longrightarrow> P)"
+lemma \<open>P \<longrightarrow> (Q \<longrightarrow> P)\<close>
by fast
\<comment> \<open>The combinator S\<close>
-lemma "(P \<longrightarrow> Q \<longrightarrow> R) \<longrightarrow> (P \<longrightarrow> Q) \<longrightarrow> (P \<longrightarrow> R)"
+lemma \<open>(P \<longrightarrow> Q \<longrightarrow> R) \<longrightarrow> (P \<longrightarrow> Q) \<longrightarrow> (P \<longrightarrow> R)\<close>
by fast
\<comment> \<open>Converse is classical\<close>
-lemma "(P \<longrightarrow> Q) \<or> (P \<longrightarrow> R) \<longrightarrow> (P \<longrightarrow> Q \<or> R)"
+lemma \<open>(P \<longrightarrow> Q) \<or> (P \<longrightarrow> R) \<longrightarrow> (P \<longrightarrow> Q \<or> R)\<close>
by fast
-lemma "(P \<longrightarrow> Q) \<longrightarrow> (\<not> Q \<longrightarrow> \<not> P)"
+lemma \<open>(P \<longrightarrow> Q) \<longrightarrow> (\<not> Q \<longrightarrow> \<not> P)\<close>
by fast
text \<open>Schwichtenberg's examples (via T. Nipkow)\<close>
-lemma stab_imp: "(((Q \<longrightarrow> R) \<longrightarrow> R) \<longrightarrow> Q) \<longrightarrow> (((P \<longrightarrow> Q) \<longrightarrow> R) \<longrightarrow> R) \<longrightarrow> P \<longrightarrow> Q"
+lemma stab_imp: \<open>(((Q \<longrightarrow> R) \<longrightarrow> R) \<longrightarrow> Q) \<longrightarrow> (((P \<longrightarrow> Q) \<longrightarrow> R) \<longrightarrow> R) \<longrightarrow> P \<longrightarrow> Q\<close>
by fast
lemma stab_to_peirce:
- "(((P \<longrightarrow> R) \<longrightarrow> R) \<longrightarrow> P) \<longrightarrow> (((Q \<longrightarrow> R) \<longrightarrow> R) \<longrightarrow> Q)
- \<longrightarrow> ((P \<longrightarrow> Q) \<longrightarrow> P) \<longrightarrow> P"
+ \<open>(((P \<longrightarrow> R) \<longrightarrow> R) \<longrightarrow> P) \<longrightarrow> (((Q \<longrightarrow> R) \<longrightarrow> R) \<longrightarrow> Q)
+ \<longrightarrow> ((P \<longrightarrow> Q) \<longrightarrow> P) \<longrightarrow> P\<close>
by fast
lemma peirce_imp1:
- "(((Q \<longrightarrow> R) \<longrightarrow> Q) \<longrightarrow> Q)
- \<longrightarrow> (((P \<longrightarrow> Q) \<longrightarrow> R) \<longrightarrow> P \<longrightarrow> Q) \<longrightarrow> P \<longrightarrow> Q"
+ \<open>(((Q \<longrightarrow> R) \<longrightarrow> Q) \<longrightarrow> Q)
+ \<longrightarrow> (((P \<longrightarrow> Q) \<longrightarrow> R) \<longrightarrow> P \<longrightarrow> Q) \<longrightarrow> P \<longrightarrow> Q\<close>
by fast
-lemma peirce_imp2: "(((P \<longrightarrow> R) \<longrightarrow> P) \<longrightarrow> P) \<longrightarrow> ((P \<longrightarrow> Q \<longrightarrow> R) \<longrightarrow> P) \<longrightarrow> P"
+lemma peirce_imp2: \<open>(((P \<longrightarrow> R) \<longrightarrow> P) \<longrightarrow> P) \<longrightarrow> ((P \<longrightarrow> Q \<longrightarrow> R) \<longrightarrow> P) \<longrightarrow> P\<close>
by fast
-lemma mints: "((((P \<longrightarrow> Q) \<longrightarrow> P) \<longrightarrow> P) \<longrightarrow> Q) \<longrightarrow> Q"
+lemma mints: \<open>((((P \<longrightarrow> Q) \<longrightarrow> P) \<longrightarrow> P) \<longrightarrow> Q) \<longrightarrow> Q\<close>
by fast
-lemma mints_solovev: "(P \<longrightarrow> (Q \<longrightarrow> R) \<longrightarrow> Q) \<longrightarrow> ((P \<longrightarrow> Q) \<longrightarrow> R) \<longrightarrow> R"
+lemma mints_solovev: \<open>(P \<longrightarrow> (Q \<longrightarrow> R) \<longrightarrow> Q) \<longrightarrow> ((P \<longrightarrow> Q) \<longrightarrow> R) \<longrightarrow> R\<close>
by fast
lemma tatsuta:
- "(((P7 \<longrightarrow> P1) \<longrightarrow> P10) \<longrightarrow> P4 \<longrightarrow> P5)
+ \<open>(((P7 \<longrightarrow> P1) \<longrightarrow> P10) \<longrightarrow> P4 \<longrightarrow> P5)
\<longrightarrow> (((P8 \<longrightarrow> P2) \<longrightarrow> P9) \<longrightarrow> P3 \<longrightarrow> P10)
\<longrightarrow> (P1 \<longrightarrow> P8) \<longrightarrow> P6 \<longrightarrow> P7
\<longrightarrow> (((P3 \<longrightarrow> P2) \<longrightarrow> P9) \<longrightarrow> P4)
- \<longrightarrow> (P1 \<longrightarrow> P3) \<longrightarrow> (((P6 \<longrightarrow> P1) \<longrightarrow> P2) \<longrightarrow> P9) \<longrightarrow> P5"
+ \<longrightarrow> (P1 \<longrightarrow> P3) \<longrightarrow> (((P6 \<longrightarrow> P1) \<longrightarrow> P2) \<longrightarrow> P9) \<longrightarrow> P5\<close>
by fast
lemma tatsuta1:
- "(((P8 \<longrightarrow> P2) \<longrightarrow> P9) \<longrightarrow> P3 \<longrightarrow> P10)
+ \<open>(((P8 \<longrightarrow> P2) \<longrightarrow> P9) \<longrightarrow> P3 \<longrightarrow> P10)
\<longrightarrow> (((P3 \<longrightarrow> P2) \<longrightarrow> P9) \<longrightarrow> P4)
\<longrightarrow> (((P6 \<longrightarrow> P1) \<longrightarrow> P2) \<longrightarrow> P9)
\<longrightarrow> (((P7 \<longrightarrow> P1) \<longrightarrow> P10) \<longrightarrow> P4 \<longrightarrow> P5)
- \<longrightarrow> (P1 \<longrightarrow> P3) \<longrightarrow> (P1 \<longrightarrow> P8) \<longrightarrow> P6 \<longrightarrow> P7 \<longrightarrow> P5"
+ \<longrightarrow> (P1 \<longrightarrow> P3) \<longrightarrow> (P1 \<longrightarrow> P8) \<longrightarrow> P6 \<longrightarrow> P7 \<longrightarrow> P5\<close>
by fast
end
--- a/src/FOL/ex/Propositional_Int.thy Thu Jan 03 21:48:05 2019 +0100
+++ b/src/FOL/ex/Propositional_Int.thy Thu Jan 03 22:19:19 2019 +0100
@@ -11,110 +11,110 @@
text \<open>commutative laws of \<open>\<and>\<close> and \<open>\<or>\<close>\<close>
-lemma "P \<and> Q \<longrightarrow> Q \<and> P"
+lemma \<open>P \<and> Q \<longrightarrow> Q \<and> P\<close>
by (tactic "IntPr.fast_tac @{context} 1")
-lemma "P \<or> Q \<longrightarrow> Q \<or> P"
+lemma \<open>P \<or> Q \<longrightarrow> Q \<or> P\<close>
by (tactic "IntPr.fast_tac @{context} 1")
text \<open>associative laws of \<open>\<and>\<close> and \<open>\<or>\<close>\<close>
-lemma "(P \<and> Q) \<and> R \<longrightarrow> P \<and> (Q \<and> R)"
+lemma \<open>(P \<and> Q) \<and> R \<longrightarrow> P \<and> (Q \<and> R)\<close>
by (tactic "IntPr.fast_tac @{context} 1")
-lemma "(P \<or> Q) \<or> R \<longrightarrow> P \<or> (Q \<or> R)"
+lemma \<open>(P \<or> Q) \<or> R \<longrightarrow> P \<or> (Q \<or> R)\<close>
by (tactic "IntPr.fast_tac @{context} 1")
text \<open>distributive laws of \<open>\<and>\<close> and \<open>\<or>\<close>\<close>
-lemma "(P \<and> Q) \<or> R \<longrightarrow> (P \<or> R) \<and> (Q \<or> R)"
+lemma \<open>(P \<and> Q) \<or> R \<longrightarrow> (P \<or> R) \<and> (Q \<or> R)\<close>
by (tactic "IntPr.fast_tac @{context} 1")
-lemma "(P \<or> R) \<and> (Q \<or> R) \<longrightarrow> (P \<and> Q) \<or> R"
+lemma \<open>(P \<or> R) \<and> (Q \<or> R) \<longrightarrow> (P \<and> Q) \<or> R\<close>
by (tactic "IntPr.fast_tac @{context} 1")
-lemma "(P \<or> Q) \<and> R \<longrightarrow> (P \<and> R) \<or> (Q \<and> R)"
+lemma \<open>(P \<or> Q) \<and> R \<longrightarrow> (P \<and> R) \<or> (Q \<and> R)\<close>
by (tactic "IntPr.fast_tac @{context} 1")
-lemma "(P \<and> R) \<or> (Q \<and> R) \<longrightarrow> (P \<or> Q) \<and> R"
+lemma \<open>(P \<and> R) \<or> (Q \<and> R) \<longrightarrow> (P \<or> Q) \<and> R\<close>
by (tactic "IntPr.fast_tac @{context} 1")
text \<open>Laws involving implication\<close>
-lemma "(P \<longrightarrow> R) \<and> (Q \<longrightarrow> R) \<longleftrightarrow> (P \<or> Q \<longrightarrow> R)"
+lemma \<open>(P \<longrightarrow> R) \<and> (Q \<longrightarrow> R) \<longleftrightarrow> (P \<or> Q \<longrightarrow> R)\<close>
by (tactic "IntPr.fast_tac @{context} 1")
-lemma "(P \<and> Q \<longrightarrow> R) \<longleftrightarrow> (P \<longrightarrow> (Q \<longrightarrow> R))"
+lemma \<open>(P \<and> Q \<longrightarrow> R) \<longleftrightarrow> (P \<longrightarrow> (Q \<longrightarrow> R))\<close>
by (tactic "IntPr.fast_tac @{context} 1")
-lemma "((P \<longrightarrow> R) \<longrightarrow> R) \<longrightarrow> ((Q \<longrightarrow> R) \<longrightarrow> R) \<longrightarrow> (P \<and> Q \<longrightarrow> R) \<longrightarrow> R"
+lemma \<open>((P \<longrightarrow> R) \<longrightarrow> R) \<longrightarrow> ((Q \<longrightarrow> R) \<longrightarrow> R) \<longrightarrow> (P \<and> Q \<longrightarrow> R) \<longrightarrow> R\<close>
by (tactic "IntPr.fast_tac @{context} 1")
-lemma "\<not> (P \<longrightarrow> R) \<longrightarrow> \<not> (Q \<longrightarrow> R) \<longrightarrow> \<not> (P \<and> Q \<longrightarrow> R)"
+lemma \<open>\<not> (P \<longrightarrow> R) \<longrightarrow> \<not> (Q \<longrightarrow> R) \<longrightarrow> \<not> (P \<and> Q \<longrightarrow> R)\<close>
by (tactic "IntPr.fast_tac @{context} 1")
-lemma "(P \<longrightarrow> Q \<and> R) \<longleftrightarrow> (P \<longrightarrow> Q) \<and> (P \<longrightarrow> R)"
+lemma \<open>(P \<longrightarrow> Q \<and> R) \<longleftrightarrow> (P \<longrightarrow> Q) \<and> (P \<longrightarrow> R)\<close>
by (tactic "IntPr.fast_tac @{context} 1")
text \<open>Propositions-as-types\<close>
\<comment> \<open>The combinator K\<close>
-lemma "P \<longrightarrow> (Q \<longrightarrow> P)"
+lemma \<open>P \<longrightarrow> (Q \<longrightarrow> P)\<close>
by (tactic "IntPr.fast_tac @{context} 1")
\<comment> \<open>The combinator S\<close>
-lemma "(P \<longrightarrow> Q \<longrightarrow> R) \<longrightarrow> (P \<longrightarrow> Q) \<longrightarrow> (P \<longrightarrow> R)"
+lemma \<open>(P \<longrightarrow> Q \<longrightarrow> R) \<longrightarrow> (P \<longrightarrow> Q) \<longrightarrow> (P \<longrightarrow> R)\<close>
by (tactic "IntPr.fast_tac @{context} 1")
\<comment> \<open>Converse is classical\<close>
-lemma "(P \<longrightarrow> Q) \<or> (P \<longrightarrow> R) \<longrightarrow> (P \<longrightarrow> Q \<or> R)"
+lemma \<open>(P \<longrightarrow> Q) \<or> (P \<longrightarrow> R) \<longrightarrow> (P \<longrightarrow> Q \<or> R)\<close>
by (tactic "IntPr.fast_tac @{context} 1")
-lemma "(P \<longrightarrow> Q) \<longrightarrow> (\<not> Q \<longrightarrow> \<not> P)"
+lemma \<open>(P \<longrightarrow> Q) \<longrightarrow> (\<not> Q \<longrightarrow> \<not> P)\<close>
by (tactic "IntPr.fast_tac @{context} 1")
text \<open>Schwichtenberg's examples (via T. Nipkow)\<close>
-lemma stab_imp: "(((Q \<longrightarrow> R) \<longrightarrow> R) \<longrightarrow> Q) \<longrightarrow> (((P \<longrightarrow> Q) \<longrightarrow> R) \<longrightarrow> R) \<longrightarrow> P \<longrightarrow> Q"
+lemma stab_imp: \<open>(((Q \<longrightarrow> R) \<longrightarrow> R) \<longrightarrow> Q) \<longrightarrow> (((P \<longrightarrow> Q) \<longrightarrow> R) \<longrightarrow> R) \<longrightarrow> P \<longrightarrow> Q\<close>
by (tactic "IntPr.fast_tac @{context} 1")
lemma stab_to_peirce:
- "(((P \<longrightarrow> R) \<longrightarrow> R) \<longrightarrow> P) \<longrightarrow> (((Q \<longrightarrow> R) \<longrightarrow> R) \<longrightarrow> Q)
- \<longrightarrow> ((P \<longrightarrow> Q) \<longrightarrow> P) \<longrightarrow> P"
+ \<open>(((P \<longrightarrow> R) \<longrightarrow> R) \<longrightarrow> P) \<longrightarrow> (((Q \<longrightarrow> R) \<longrightarrow> R) \<longrightarrow> Q)
+ \<longrightarrow> ((P \<longrightarrow> Q) \<longrightarrow> P) \<longrightarrow> P\<close>
by (tactic "IntPr.fast_tac @{context} 1")
lemma peirce_imp1:
- "(((Q \<longrightarrow> R) \<longrightarrow> Q) \<longrightarrow> Q)
- \<longrightarrow> (((P \<longrightarrow> Q) \<longrightarrow> R) \<longrightarrow> P \<longrightarrow> Q) \<longrightarrow> P \<longrightarrow> Q"
+ \<open>(((Q \<longrightarrow> R) \<longrightarrow> Q) \<longrightarrow> Q)
+ \<longrightarrow> (((P \<longrightarrow> Q) \<longrightarrow> R) \<longrightarrow> P \<longrightarrow> Q) \<longrightarrow> P \<longrightarrow> Q\<close>
by (tactic "IntPr.fast_tac @{context} 1")
-lemma peirce_imp2: "(((P \<longrightarrow> R) \<longrightarrow> P) \<longrightarrow> P) \<longrightarrow> ((P \<longrightarrow> Q \<longrightarrow> R) \<longrightarrow> P) \<longrightarrow> P"
+lemma peirce_imp2: \<open>(((P \<longrightarrow> R) \<longrightarrow> P) \<longrightarrow> P) \<longrightarrow> ((P \<longrightarrow> Q \<longrightarrow> R) \<longrightarrow> P) \<longrightarrow> P\<close>
by (tactic "IntPr.fast_tac @{context} 1")
-lemma mints: "((((P \<longrightarrow> Q) \<longrightarrow> P) \<longrightarrow> P) \<longrightarrow> Q) \<longrightarrow> Q"
+lemma mints: \<open>((((P \<longrightarrow> Q) \<longrightarrow> P) \<longrightarrow> P) \<longrightarrow> Q) \<longrightarrow> Q\<close>
by (tactic "IntPr.fast_tac @{context} 1")
-lemma mints_solovev: "(P \<longrightarrow> (Q \<longrightarrow> R) \<longrightarrow> Q) \<longrightarrow> ((P \<longrightarrow> Q) \<longrightarrow> R) \<longrightarrow> R"
+lemma mints_solovev: \<open>(P \<longrightarrow> (Q \<longrightarrow> R) \<longrightarrow> Q) \<longrightarrow> ((P \<longrightarrow> Q) \<longrightarrow> R) \<longrightarrow> R\<close>
by (tactic "IntPr.fast_tac @{context} 1")
lemma tatsuta:
- "(((P7 \<longrightarrow> P1) \<longrightarrow> P10) \<longrightarrow> P4 \<longrightarrow> P5)
+ \<open>(((P7 \<longrightarrow> P1) \<longrightarrow> P10) \<longrightarrow> P4 \<longrightarrow> P5)
\<longrightarrow> (((P8 \<longrightarrow> P2) \<longrightarrow> P9) \<longrightarrow> P3 \<longrightarrow> P10)
\<longrightarrow> (P1 \<longrightarrow> P8) \<longrightarrow> P6 \<longrightarrow> P7
\<longrightarrow> (((P3 \<longrightarrow> P2) \<longrightarrow> P9) \<longrightarrow> P4)
- \<longrightarrow> (P1 \<longrightarrow> P3) \<longrightarrow> (((P6 \<longrightarrow> P1) \<longrightarrow> P2) \<longrightarrow> P9) \<longrightarrow> P5"
+ \<longrightarrow> (P1 \<longrightarrow> P3) \<longrightarrow> (((P6 \<longrightarrow> P1) \<longrightarrow> P2) \<longrightarrow> P9) \<longrightarrow> P5\<close>
by (tactic "IntPr.fast_tac @{context} 1")
lemma tatsuta1:
- "(((P8 \<longrightarrow> P2) \<longrightarrow> P9) \<longrightarrow> P3 \<longrightarrow> P10)
+ \<open>(((P8 \<longrightarrow> P2) \<longrightarrow> P9) \<longrightarrow> P3 \<longrightarrow> P10)
\<longrightarrow> (((P3 \<longrightarrow> P2) \<longrightarrow> P9) \<longrightarrow> P4)
\<longrightarrow> (((P6 \<longrightarrow> P1) \<longrightarrow> P2) \<longrightarrow> P9)
\<longrightarrow> (((P7 \<longrightarrow> P1) \<longrightarrow> P10) \<longrightarrow> P4 \<longrightarrow> P5)
- \<longrightarrow> (P1 \<longrightarrow> P3) \<longrightarrow> (P1 \<longrightarrow> P8) \<longrightarrow> P6 \<longrightarrow> P7 \<longrightarrow> P5"
+ \<longrightarrow> (P1 \<longrightarrow> P3) \<longrightarrow> (P1 \<longrightarrow> P8) \<longrightarrow> P6 \<longrightarrow> P7 \<longrightarrow> P5\<close>
by (tactic "IntPr.fast_tac @{context} 1")
end
--- a/src/FOL/ex/Quantifiers_Cla.thy Thu Jan 03 21:48:05 2019 +0100
+++ b/src/FOL/ex/Quantifiers_Cla.thy Thu Jan 03 22:19:19 2019 +0100
@@ -9,92 +9,92 @@
imports FOL
begin
-lemma "(\<forall>x y. P(x,y)) \<longrightarrow> (\<forall>y x. P(x,y))"
+lemma \<open>(\<forall>x y. P(x,y)) \<longrightarrow> (\<forall>y x. P(x,y))\<close>
by fast
-lemma "(\<exists>x y. P(x,y)) \<longrightarrow> (\<exists>y x. P(x,y))"
+lemma \<open>(\<exists>x y. P(x,y)) \<longrightarrow> (\<exists>y x. P(x,y))\<close>
by fast
text \<open>Converse is false.\<close>
-lemma "(\<forall>x. P(x)) \<or> (\<forall>x. Q(x)) \<longrightarrow> (\<forall>x. P(x) \<or> Q(x))"
+lemma \<open>(\<forall>x. P(x)) \<or> (\<forall>x. Q(x)) \<longrightarrow> (\<forall>x. P(x) \<or> Q(x))\<close>
by fast
-lemma "(\<forall>x. P \<longrightarrow> Q(x)) \<longleftrightarrow> (P \<longrightarrow> (\<forall>x. Q(x)))"
+lemma \<open>(\<forall>x. P \<longrightarrow> Q(x)) \<longleftrightarrow> (P \<longrightarrow> (\<forall>x. Q(x)))\<close>
by fast
-lemma "(\<forall>x. P(x) \<longrightarrow> Q) \<longleftrightarrow> ((\<exists>x. P(x)) \<longrightarrow> Q)"
+lemma \<open>(\<forall>x. P(x) \<longrightarrow> Q) \<longleftrightarrow> ((\<exists>x. P(x)) \<longrightarrow> Q)\<close>
by fast
text \<open>Some harder ones.\<close>
-lemma "(\<exists>x. P(x) \<or> Q(x)) \<longleftrightarrow> (\<exists>x. P(x)) \<or> (\<exists>x. Q(x))"
+lemma \<open>(\<exists>x. P(x) \<or> Q(x)) \<longleftrightarrow> (\<exists>x. P(x)) \<or> (\<exists>x. Q(x))\<close>
by fast
\<comment> \<open>Converse is false.\<close>
-lemma "(\<exists>x. P(x) \<and> Q(x)) \<longrightarrow> (\<exists>x. P(x)) \<and> (\<exists>x. Q(x))"
+lemma \<open>(\<exists>x. P(x) \<and> Q(x)) \<longrightarrow> (\<exists>x. P(x)) \<and> (\<exists>x. Q(x))\<close>
by fast
text \<open>Basic test of quantifier reasoning.\<close>
\<comment> \<open>TRUE\<close>
-lemma "(\<exists>y. \<forall>x. Q(x,y)) \<longrightarrow> (\<forall>x. \<exists>y. Q(x,y))"
+lemma \<open>(\<exists>y. \<forall>x. Q(x,y)) \<longrightarrow> (\<forall>x. \<exists>y. Q(x,y))\<close>
by fast
-lemma "(\<forall>x. Q(x)) \<longrightarrow> (\<exists>x. Q(x))"
+lemma \<open>(\<forall>x. Q(x)) \<longrightarrow> (\<exists>x. Q(x))\<close>
by fast
text \<open>The following should fail, as they are false!\<close>
-lemma "(\<forall>x. \<exists>y. Q(x,y)) \<longrightarrow> (\<exists>y. \<forall>x. Q(x,y))"
+lemma \<open>(\<forall>x. \<exists>y. Q(x,y)) \<longrightarrow> (\<exists>y. \<forall>x. Q(x,y))\<close>
apply fast?
oops
-lemma "(\<exists>x. Q(x)) \<longrightarrow> (\<forall>x. Q(x))"
+lemma \<open>(\<exists>x. Q(x)) \<longrightarrow> (\<forall>x. Q(x))\<close>
apply fast?
oops
-schematic_goal "P(?a) \<longrightarrow> (\<forall>x. P(x))"
+schematic_goal \<open>P(?a) \<longrightarrow> (\<forall>x. P(x))\<close>
apply fast?
oops
-schematic_goal "(P(?a) \<longrightarrow> (\<forall>x. Q(x))) \<longrightarrow> (\<forall>x. P(x) \<longrightarrow> Q(x))"
+schematic_goal \<open>(P(?a) \<longrightarrow> (\<forall>x. Q(x))) \<longrightarrow> (\<forall>x. P(x) \<longrightarrow> Q(x))\<close>
apply fast?
oops
text \<open>Back to things that are provable \dots\<close>
-lemma "(\<forall>x. P(x) \<longrightarrow> Q(x)) \<and> (\<exists>x. P(x)) \<longrightarrow> (\<exists>x. Q(x))"
+lemma \<open>(\<forall>x. P(x) \<longrightarrow> Q(x)) \<and> (\<exists>x. P(x)) \<longrightarrow> (\<exists>x. Q(x))\<close>
by fast
text \<open>An example of why \<open>exI\<close> should be delayed as long as possible.\<close>
-lemma "(P \<longrightarrow> (\<exists>x. Q(x))) \<and> P \<longrightarrow> (\<exists>x. Q(x))"
+lemma \<open>(P \<longrightarrow> (\<exists>x. Q(x))) \<and> P \<longrightarrow> (\<exists>x. Q(x))\<close>
by fast
-schematic_goal "(\<forall>x. P(x) \<longrightarrow> Q(f(x))) \<and> (\<forall>x. Q(x) \<longrightarrow> R(g(x))) \<and> P(d) \<longrightarrow> R(?a)"
+schematic_goal \<open>(\<forall>x. P(x) \<longrightarrow> Q(f(x))) \<and> (\<forall>x. Q(x) \<longrightarrow> R(g(x))) \<and> P(d) \<longrightarrow> R(?a)\<close>
by fast
-lemma "(\<forall>x. Q(x)) \<longrightarrow> (\<exists>x. Q(x))"
+lemma \<open>(\<forall>x. Q(x)) \<longrightarrow> (\<exists>x. Q(x))\<close>
by fast
text \<open>Some slow ones\<close>
text \<open>Principia Mathematica *11.53\<close>
-lemma "(\<forall>x y. P(x) \<longrightarrow> Q(y)) \<longleftrightarrow> ((\<exists>x. P(x)) \<longrightarrow> (\<forall>y. Q(y)))"
+lemma \<open>(\<forall>x y. P(x) \<longrightarrow> Q(y)) \<longleftrightarrow> ((\<exists>x. P(x)) \<longrightarrow> (\<forall>y. Q(y)))\<close>
by fast
(*Principia Mathematica *11.55 *)
-lemma "(\<exists>x y. P(x) \<and> Q(x,y)) \<longleftrightarrow> (\<exists>x. P(x) \<and> (\<exists>y. Q(x,y)))"
+lemma \<open>(\<exists>x y. P(x) \<and> Q(x,y)) \<longleftrightarrow> (\<exists>x. P(x) \<and> (\<exists>y. Q(x,y)))\<close>
by fast
(*Principia Mathematica *11.61 *)
-lemma "(\<exists>y. \<forall>x. P(x) \<longrightarrow> Q(x,y)) \<longrightarrow> (\<forall>x. P(x) \<longrightarrow> (\<exists>y. Q(x,y)))"
+lemma \<open>(\<exists>y. \<forall>x. P(x) \<longrightarrow> Q(x,y)) \<longrightarrow> (\<forall>x. P(x) \<longrightarrow> (\<exists>y. Q(x,y)))\<close>
by fast
end
--- a/src/FOL/ex/Quantifiers_Int.thy Thu Jan 03 21:48:05 2019 +0100
+++ b/src/FOL/ex/Quantifiers_Int.thy Thu Jan 03 22:19:19 2019 +0100
@@ -9,92 +9,92 @@
imports IFOL
begin
-lemma "(\<forall>x y. P(x,y)) \<longrightarrow> (\<forall>y x. P(x,y))"
+lemma \<open>(\<forall>x y. P(x,y)) \<longrightarrow> (\<forall>y x. P(x,y))\<close>
by (tactic "IntPr.fast_tac @{context} 1")
-lemma "(\<exists>x y. P(x,y)) \<longrightarrow> (\<exists>y x. P(x,y))"
+lemma \<open>(\<exists>x y. P(x,y)) \<longrightarrow> (\<exists>y x. P(x,y))\<close>
by (tactic "IntPr.fast_tac @{context} 1")
\<comment> \<open>Converse is false\<close>
-lemma "(\<forall>x. P(x)) \<or> (\<forall>x. Q(x)) \<longrightarrow> (\<forall>x. P(x) \<or> Q(x))"
+lemma \<open>(\<forall>x. P(x)) \<or> (\<forall>x. Q(x)) \<longrightarrow> (\<forall>x. P(x) \<or> Q(x))\<close>
by (tactic "IntPr.fast_tac @{context} 1")
-lemma "(\<forall>x. P \<longrightarrow> Q(x)) \<longleftrightarrow> (P \<longrightarrow> (\<forall>x. Q(x)))"
+lemma \<open>(\<forall>x. P \<longrightarrow> Q(x)) \<longleftrightarrow> (P \<longrightarrow> (\<forall>x. Q(x)))\<close>
by (tactic "IntPr.fast_tac @{context} 1")
-lemma "(\<forall>x. P(x) \<longrightarrow> Q) \<longleftrightarrow> ((\<exists>x. P(x)) \<longrightarrow> Q)"
+lemma \<open>(\<forall>x. P(x) \<longrightarrow> Q) \<longleftrightarrow> ((\<exists>x. P(x)) \<longrightarrow> Q)\<close>
by (tactic "IntPr.fast_tac @{context} 1")
text \<open>Some harder ones\<close>
-lemma "(\<exists>x. P(x) \<or> Q(x)) \<longleftrightarrow> (\<exists>x. P(x)) \<or> (\<exists>x. Q(x))"
+lemma \<open>(\<exists>x. P(x) \<or> Q(x)) \<longleftrightarrow> (\<exists>x. P(x)) \<or> (\<exists>x. Q(x))\<close>
by (tactic "IntPr.fast_tac @{context} 1")
\<comment> \<open>Converse is false\<close>
-lemma "(\<exists>x. P(x) \<and> Q(x)) \<longrightarrow> (\<exists>x. P(x)) \<and> (\<exists>x. Q(x))"
+lemma \<open>(\<exists>x. P(x) \<and> Q(x)) \<longrightarrow> (\<exists>x. P(x)) \<and> (\<exists>x. Q(x))\<close>
by (tactic "IntPr.fast_tac @{context} 1")
text \<open>Basic test of quantifier reasoning\<close>
\<comment> \<open>TRUE\<close>
-lemma "(\<exists>y. \<forall>x. Q(x,y)) \<longrightarrow> (\<forall>x. \<exists>y. Q(x,y))"
+lemma \<open>(\<exists>y. \<forall>x. Q(x,y)) \<longrightarrow> (\<forall>x. \<exists>y. Q(x,y))\<close>
by (tactic "IntPr.fast_tac @{context} 1")
-lemma "(\<forall>x. Q(x)) \<longrightarrow> (\<exists>x. Q(x))"
+lemma \<open>(\<forall>x. Q(x)) \<longrightarrow> (\<exists>x. Q(x))\<close>
by (tactic "IntPr.fast_tac @{context} 1")
text \<open>The following should fail, as they are false!\<close>
-lemma "(\<forall>x. \<exists>y. Q(x,y)) \<longrightarrow> (\<exists>y. \<forall>x. Q(x,y))"
+lemma \<open>(\<forall>x. \<exists>y. Q(x,y)) \<longrightarrow> (\<exists>y. \<forall>x. Q(x,y))\<close>
apply (tactic "IntPr.fast_tac @{context} 1")?
oops
-lemma "(\<exists>x. Q(x)) \<longrightarrow> (\<forall>x. Q(x))"
+lemma \<open>(\<exists>x. Q(x)) \<longrightarrow> (\<forall>x. Q(x))\<close>
apply (tactic "IntPr.fast_tac @{context} 1")?
oops
-schematic_goal "P(?a) \<longrightarrow> (\<forall>x. P(x))"
+schematic_goal \<open>P(?a) \<longrightarrow> (\<forall>x. P(x))\<close>
apply (tactic "IntPr.fast_tac @{context} 1")?
oops
-schematic_goal "(P(?a) \<longrightarrow> (\<forall>x. Q(x))) \<longrightarrow> (\<forall>x. P(x) \<longrightarrow> Q(x))"
+schematic_goal \<open>(P(?a) \<longrightarrow> (\<forall>x. Q(x))) \<longrightarrow> (\<forall>x. P(x) \<longrightarrow> Q(x))\<close>
apply (tactic "IntPr.fast_tac @{context} 1")?
oops
text \<open>Back to things that are provable \dots\<close>
-lemma "(\<forall>x. P(x) \<longrightarrow> Q(x)) \<and> (\<exists>x. P(x)) \<longrightarrow> (\<exists>x. Q(x))"
+lemma \<open>(\<forall>x. P(x) \<longrightarrow> Q(x)) \<and> (\<exists>x. P(x)) \<longrightarrow> (\<exists>x. Q(x))\<close>
by (tactic "IntPr.fast_tac @{context} 1")
\<comment> \<open>An example of why exI should be delayed as long as possible\<close>
-lemma "(P \<longrightarrow> (\<exists>x. Q(x))) \<and> P \<longrightarrow> (\<exists>x. Q(x))"
+lemma \<open>(P \<longrightarrow> (\<exists>x. Q(x))) \<and> P \<longrightarrow> (\<exists>x. Q(x))\<close>
by (tactic "IntPr.fast_tac @{context} 1")
-schematic_goal "(\<forall>x. P(x) \<longrightarrow> Q(f(x))) \<and> (\<forall>x. Q(x) \<longrightarrow> R(g(x))) \<and> P(d) \<longrightarrow> R(?a)"
+schematic_goal \<open>(\<forall>x. P(x) \<longrightarrow> Q(f(x))) \<and> (\<forall>x. Q(x) \<longrightarrow> R(g(x))) \<and> P(d) \<longrightarrow> R(?a)\<close>
by (tactic "IntPr.fast_tac @{context} 1")
-lemma "(\<forall>x. Q(x)) \<longrightarrow> (\<exists>x. Q(x))"
+lemma \<open>(\<forall>x. Q(x)) \<longrightarrow> (\<exists>x. Q(x))\<close>
by (tactic "IntPr.fast_tac @{context} 1")
text \<open>Some slow ones\<close>
\<comment> \<open>Principia Mathematica *11.53\<close>
-lemma "(\<forall>x y. P(x) \<longrightarrow> Q(y)) \<longleftrightarrow> ((\<exists>x. P(x)) \<longrightarrow> (\<forall>y. Q(y)))"
+lemma \<open>(\<forall>x y. P(x) \<longrightarrow> Q(y)) \<longleftrightarrow> ((\<exists>x. P(x)) \<longrightarrow> (\<forall>y. Q(y)))\<close>
by (tactic "IntPr.fast_tac @{context} 1")
(*Principia Mathematica *11.55 *)
-lemma "(\<exists>x y. P(x) \<and> Q(x,y)) \<longleftrightarrow> (\<exists>x. P(x) \<and> (\<exists>y. Q(x,y)))"
+lemma \<open>(\<exists>x y. P(x) \<and> Q(x,y)) \<longleftrightarrow> (\<exists>x. P(x) \<and> (\<exists>y. Q(x,y)))\<close>
by (tactic "IntPr.fast_tac @{context} 1")
(*Principia Mathematica *11.61 *)
-lemma "(\<exists>y. \<forall>x. P(x) \<longrightarrow> Q(x,y)) \<longrightarrow> (\<forall>x. P(x) \<longrightarrow> (\<exists>y. Q(x,y)))"
+lemma \<open>(\<exists>y. \<forall>x. P(x) \<longrightarrow> Q(x,y)) \<longrightarrow> (\<forall>x. P(x) \<longrightarrow> (\<exists>y. Q(x,y)))\<close>
by (tactic "IntPr.fast_tac @{context} 1")
end
--- a/src/Tools/induct.ML Thu Jan 03 21:48:05 2019 +0100
+++ b/src/Tools/induct.ML Thu Jan 03 22:19:19 2019 +0100
@@ -172,7 +172,7 @@
val rearrange_eqs_simproc =
Simplifier.make_simproc @{context} "rearrange_eqs"
- {lhss = [@{term "Pure.all (t :: 'a::{} \<Rightarrow> prop)"}],
+ {lhss = [@{term \<open>Pure.all (t :: 'a::{} \<Rightarrow> prop)\<close>}],
proc = fn _ => fn ctxt => fn ct => mk_swap_rrule ctxt ct};