introduction of "by" and a few examples of SOME
authorpaulson
Wed, 10 Jan 2001 11:06:07 +0100
changeset 10843 f2e4e383dbca
parent 10842 4649e5e3905d
child 10844 0c0e7de7e9c5
introduction of "by" and a few examples of SOME
doc-src/TutorialI/Rules/Basic.thy
--- a/doc-src/TutorialI/Rules/Basic.thy	Wed Jan 10 11:05:27 2001 +0100
+++ b/doc-src/TutorialI/Rules/Basic.thy	Wed Jan 10 11:06:07 2001 +0100
@@ -36,6 +36,18 @@
  apply assumption
 done
 
+text {*NEW
+by eliminates uses of assumption and done
+*}
+
+lemma imp_uncurry: "P \<longrightarrow> Q \<longrightarrow> R \<Longrightarrow> P \<and> Q \<longrightarrow> R"
+apply (rule impI)
+apply (erule conjE)
+apply (drule mp)
+ apply assumption
+by (drule mp)
+
+
 text {*
 substitution
 
@@ -44,9 +56,7 @@
 *};
 
 lemma "\<lbrakk> x = f x; P(f x) \<rbrakk> \<Longrightarrow> P x"
-apply (erule ssubst)
-apply assumption
-done
+by (erule ssubst)
 
 text {*
 also provable by simp (re-orients)
@@ -97,12 +107,25 @@
 apply (erule ssubst, assumption)
 done
 
+text{*
+or better still NEW
+*}
+
 lemma "\<lbrakk> x = f x; P (f x) (f x) x \<rbrakk> \<Longrightarrow> P x x x"
-apply (erule_tac P="\<lambda>u. P u u x" in ssubst);
-apply assumption
+by (erule ssubst)
+
+
+text{*NEW*}
+lemma "\<lbrakk> x = f x; P (f x) (f x) x \<rbrakk> \<Longrightarrow> P x x x"
+apply (erule_tac P="\<lambda>u. P u u x" in ssubst)
+apply (assumption)
 done
 
 
+lemma "\<lbrakk> x = f x; P (f x) (f x) x \<rbrakk> \<Longrightarrow> P x x x"
+by (erule_tac P="\<lambda>u. P u u x" in ssubst)
+
+
 text {*
 negation
 
@@ -144,8 +167,9 @@
 {\isasymlbrakk}{\isasymnot}\ {\isacharparenleft}P\ {\isasymlongrightarrow}\ Q{\isacharparenright}{\isacharsemicolon}\ {\isasymnot}\ {\isacharparenleft}R\ {\isasymlongrightarrow}\ Q{\isacharparenright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ R\isanewline
 \ {\isadigit{1}}{\isachardot}\ {\isasymlbrakk}{\isasymnot}\ {\isacharparenleft}P\ {\isasymlongrightarrow}\ Q{\isacharparenright}{\isacharsemicolon}\ {\isasymnot}\ R{\isacharsemicolon}\ R{\isasymrbrakk}\ {\isasymLongrightarrow}\ Q
 *}
-apply (erule notE, assumption)
-done
+by (erule notE)
+text{*NEW*}
+
 
 
 lemma "(P \<or> Q) \<and> R \<Longrightarrow> P \<or> Q \<and> R"
@@ -169,8 +193,8 @@
 \ {\isadigit{1}}{\isachardot}\ {\isasymlbrakk}{\isasymnot}\ {\isacharparenleft}Q\ {\isasymand}\ R{\isacharparenright}{\isacharsemicolon}\ R{\isacharsemicolon}\ Q{\isasymrbrakk}\ {\isasymLongrightarrow}\ P
 *}
 
-apply (erule contrapos_np, rule conjI)
-txt{*
+by (erule contrapos_np, rule conjI)
+text{*NEW
 proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ {\isadigit{6}}\isanewline
 \isanewline
 goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline
@@ -179,35 +203,27 @@
 \ {\isadigit{2}}{\isachardot}\ {\isasymlbrakk}R{\isacharsemicolon}\ Q{\isacharsemicolon}\ {\isasymnot}\ P{\isasymrbrakk}\ {\isasymLongrightarrow}\ R
 *}
 
-  apply assumption
- apply assumption
-done
-
-
 
 text{*Quantifiers*}
 
 lemma "\<forall>x. P x \<longrightarrow> P x"
 apply (rule allI)
-apply (rule impI)
-apply assumption
-done
+by (rule impI)
+text{*NEW*}
 
 lemma "(\<forall>x. P \<longrightarrow> Q x) \<Longrightarrow> P \<longrightarrow> (\<forall>x. Q x)"
-apply (rule impI)
-apply (rule allI)
+apply (rule impI, rule allI)
 apply (drule spec)
-apply (drule mp)
-  apply assumption
- apply assumption
-done
+by (drule mp)
+text{*NEW*}
 
-lemma "\<lbrakk>\<forall>x. P x \<longrightarrow> P (f x); P a\<rbrakk> \<Longrightarrow> P(f (f a))"
+lemma "\<lbrakk>\<forall>x. P x \<longrightarrow> P (h x); P a\<rbrakk> \<Longrightarrow> P(h (h a))"
 apply (frule spec)
 apply (drule mp, assumption)
 apply (drule spec)
-apply (drule mp, assumption, assumption)
-done
+by (drule mp)
+text{*NEW*}
+
 
 text
 {*
@@ -229,6 +245,86 @@
 lemma "\<lbrakk>\<forall>x. P x \<longrightarrow> P (f x); P a\<rbrakk> \<Longrightarrow> P(f (f a))"
 by blast
 
+text{*NEW
+Hilbert-epsilon theorems*}
+
+text{*
+@{thm[display] some_equality[no_vars]}
+\rulename{some_equality}
+
+@{thm[display] someI[no_vars]}
+\rulename{someI}
+
+@{thm[display] someI2[no_vars]}
+\rulename{someI2}
+
+needed for examples
+
+@{thm[display] inv_def[no_vars]}
+\rulename{inv_def}
+
+@{thm[display] Least_def[no_vars]}
+\rulename{Least_def}
+
+@{thm[display] order_antisym[no_vars]}
+\rulename{order_antisym}
+*}
+
+
+lemma "inv Suc (Suc x) = x"
+by (simp add: inv_def)
+
+text{*but we know nothing about inv Suc 0*}
+
+theorem Least_equality:
+     "\<lbrakk> P (k::nat);  \<forall>x. P x \<longrightarrow> k \<le> x \<rbrakk> \<Longrightarrow> (LEAST x. P(x)) = k"
+apply (simp add: Least_def)
+ 
+txt{*omit maybe?
+@{subgoals[display,indent=0,margin=65]}
+*};
+   
+apply (rule some_equality)
+
+txt{*
+@{subgoals[display,indent=0,margin=65]}
+
+first subgoal is existence; second is uniqueness
+*};
+by (auto intro: order_antisym)
+
+
+theorem axiom_of_choice:
+     "(\<forall>x. \<exists> y. P x y) \<Longrightarrow> \<exists>f. \<forall>x. P x (f x)"
+apply (rule exI, rule allI)
+
+txt{*
+@{subgoals[display,indent=0,margin=65]}
+
+state after intro rules
+*};
+apply (drule spec, erule exE)
+
+txt{*
+@{subgoals[display,indent=0,margin=65]}
+
+applying @text{someI} automatically instantiates
+@{term f} to @{term "\<lambda>x. SOME y. P x y"}
+*};
+
+by (rule someI)
+
+(*both can be done by blast, which however hasn't been introduced yet*)
+lemma "[| P (k::nat);  \<forall>x. P x \<longrightarrow> k \<le> x |] ==> (LEAST x. P(x)) = k";
+apply (simp add: Least_def)
+by (blast intro: some_equality order_antisym);
+
+theorem axiom_of_choice: "(\<forall>x. \<exists> y. P x y) \<Longrightarrow> \<exists>f. \<forall>x. P x (f x)"
+apply (rule exI [of _  "\<lambda>x. SOME y. P x y"])
+by (blast intro: someI);
+
+text{*end of NEW material*}
+
 lemma "(\<exists>x. P x) \<or> (\<exists>x. Q x) \<Longrightarrow> \<exists>x. P x \<or> Q x"
 apply elim
  apply intro