doc-src/TutorialI/Rules/Basic.thy
changeset 10843 f2e4e383dbca
parent 10341 6eb91805a012
child 10957 2a4a50e7ddf2
equal deleted inserted replaced
10842:4649e5e3905d 10843:f2e4e383dbca
    34 apply (drule mp)
    34 apply (drule mp)
    35   apply assumption
    35   apply assumption
    36  apply assumption
    36  apply assumption
    37 done
    37 done
    38 
    38 
       
    39 text {*NEW
       
    40 by eliminates uses of assumption and done
       
    41 *}
       
    42 
       
    43 lemma imp_uncurry: "P \<longrightarrow> Q \<longrightarrow> R \<Longrightarrow> P \<and> Q \<longrightarrow> R"
       
    44 apply (rule impI)
       
    45 apply (erule conjE)
       
    46 apply (drule mp)
       
    47  apply assumption
       
    48 by (drule mp)
       
    49 
       
    50 
    39 text {*
    51 text {*
    40 substitution
    52 substitution
    41 
    53 
    42 @{thm[display] ssubst}
    54 @{thm[display] ssubst}
    43 \rulename{ssubst}
    55 \rulename{ssubst}
    44 *};
    56 *};
    45 
    57 
    46 lemma "\<lbrakk> x = f x; P(f x) \<rbrakk> \<Longrightarrow> P x"
    58 lemma "\<lbrakk> x = f x; P(f x) \<rbrakk> \<Longrightarrow> P x"
    47 apply (erule ssubst)
    59 by (erule ssubst)
    48 apply assumption
       
    49 done
       
    50 
    60 
    51 text {*
    61 text {*
    52 also provable by simp (re-orients)
    62 also provable by simp (re-orients)
    53 *};
    63 *};
    54 
    64 
    95 
   105 
    96 lemma "\<lbrakk> x = f x; P (f x) (f x) x \<rbrakk> \<Longrightarrow> P x x x"
   106 lemma "\<lbrakk> x = f x; P (f x) (f x) x \<rbrakk> \<Longrightarrow> P x x x"
    97 apply (erule ssubst, assumption)
   107 apply (erule ssubst, assumption)
    98 done
   108 done
    99 
   109 
   100 lemma "\<lbrakk> x = f x; P (f x) (f x) x \<rbrakk> \<Longrightarrow> P x x x"
   110 text{*
   101 apply (erule_tac P="\<lambda>u. P u u x" in ssubst);
   111 or better still NEW
   102 apply assumption
   112 *}
   103 done
   113 
       
   114 lemma "\<lbrakk> x = f x; P (f x) (f x) x \<rbrakk> \<Longrightarrow> P x x x"
       
   115 by (erule ssubst)
       
   116 
       
   117 
       
   118 text{*NEW*}
       
   119 lemma "\<lbrakk> x = f x; P (f x) (f x) x \<rbrakk> \<Longrightarrow> P x x x"
       
   120 apply (erule_tac P="\<lambda>u. P u u x" in ssubst)
       
   121 apply (assumption)
       
   122 done
       
   123 
       
   124 
       
   125 lemma "\<lbrakk> x = f x; P (f x) (f x) x \<rbrakk> \<Longrightarrow> P x x x"
       
   126 by (erule_tac P="\<lambda>u. P u u x" in ssubst)
   104 
   127 
   105 
   128 
   106 text {*
   129 text {*
   107 negation
   130 negation
   108 
   131 
   142 \isanewline
   165 \isanewline
   143 goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline
   166 goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline
   144 {\isasymlbrakk}{\isasymnot}\ {\isacharparenleft}P\ {\isasymlongrightarrow}\ Q{\isacharparenright}{\isacharsemicolon}\ {\isasymnot}\ {\isacharparenleft}R\ {\isasymlongrightarrow}\ Q{\isacharparenright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ R\isanewline
   167 {\isasymlbrakk}{\isasymnot}\ {\isacharparenleft}P\ {\isasymlongrightarrow}\ Q{\isacharparenright}{\isacharsemicolon}\ {\isasymnot}\ {\isacharparenleft}R\ {\isasymlongrightarrow}\ Q{\isacharparenright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ R\isanewline
   145 \ {\isadigit{1}}{\isachardot}\ {\isasymlbrakk}{\isasymnot}\ {\isacharparenleft}P\ {\isasymlongrightarrow}\ Q{\isacharparenright}{\isacharsemicolon}\ {\isasymnot}\ R{\isacharsemicolon}\ R{\isasymrbrakk}\ {\isasymLongrightarrow}\ Q
   168 \ {\isadigit{1}}{\isachardot}\ {\isasymlbrakk}{\isasymnot}\ {\isacharparenleft}P\ {\isasymlongrightarrow}\ Q{\isacharparenright}{\isacharsemicolon}\ {\isasymnot}\ R{\isacharsemicolon}\ R{\isasymrbrakk}\ {\isasymLongrightarrow}\ Q
   146 *}
   169 *}
   147 apply (erule notE, assumption)
   170 by (erule notE)
   148 done
   171 text{*NEW*}
       
   172 
   149 
   173 
   150 
   174 
   151 lemma "(P \<or> Q) \<and> R \<Longrightarrow> P \<or> Q \<and> R"
   175 lemma "(P \<or> Q) \<and> R \<Longrightarrow> P \<or> Q \<and> R"
   152 apply intro
   176 apply intro
   153 txt{*
   177 txt{*
   167 goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline
   191 goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline
   168 {\isacharparenleft}P\ {\isasymor}\ Q{\isacharparenright}\ {\isasymand}\ R\ {\isasymLongrightarrow}\ P\ {\isasymor}\ Q\ {\isasymand}\ R\isanewline
   192 {\isacharparenleft}P\ {\isasymor}\ Q{\isacharparenright}\ {\isasymand}\ R\ {\isasymLongrightarrow}\ P\ {\isasymor}\ Q\ {\isasymand}\ R\isanewline
   169 \ {\isadigit{1}}{\isachardot}\ {\isasymlbrakk}{\isasymnot}\ {\isacharparenleft}Q\ {\isasymand}\ R{\isacharparenright}{\isacharsemicolon}\ R{\isacharsemicolon}\ Q{\isasymrbrakk}\ {\isasymLongrightarrow}\ P
   193 \ {\isadigit{1}}{\isachardot}\ {\isasymlbrakk}{\isasymnot}\ {\isacharparenleft}Q\ {\isasymand}\ R{\isacharparenright}{\isacharsemicolon}\ R{\isacharsemicolon}\ Q{\isasymrbrakk}\ {\isasymLongrightarrow}\ P
   170 *}
   194 *}
   171 
   195 
   172 apply (erule contrapos_np, rule conjI)
   196 by (erule contrapos_np, rule conjI)
   173 txt{*
   197 text{*NEW
   174 proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ {\isadigit{6}}\isanewline
   198 proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ {\isadigit{6}}\isanewline
   175 \isanewline
   199 \isanewline
   176 goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline
   200 goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline
   177 {\isacharparenleft}P\ {\isasymor}\ Q{\isacharparenright}\ {\isasymand}\ R\ {\isasymLongrightarrow}\ P\ {\isasymor}\ Q\ {\isasymand}\ R\isanewline
   201 {\isacharparenleft}P\ {\isasymor}\ Q{\isacharparenright}\ {\isasymand}\ R\ {\isasymLongrightarrow}\ P\ {\isasymor}\ Q\ {\isasymand}\ R\isanewline
   178 \ {\isadigit{1}}{\isachardot}\ {\isasymlbrakk}R{\isacharsemicolon}\ Q{\isacharsemicolon}\ {\isasymnot}\ P{\isasymrbrakk}\ {\isasymLongrightarrow}\ Q\isanewline
   202 \ {\isadigit{1}}{\isachardot}\ {\isasymlbrakk}R{\isacharsemicolon}\ Q{\isacharsemicolon}\ {\isasymnot}\ P{\isasymrbrakk}\ {\isasymLongrightarrow}\ Q\isanewline
   179 \ {\isadigit{2}}{\isachardot}\ {\isasymlbrakk}R{\isacharsemicolon}\ Q{\isacharsemicolon}\ {\isasymnot}\ P{\isasymrbrakk}\ {\isasymLongrightarrow}\ R
   203 \ {\isadigit{2}}{\isachardot}\ {\isasymlbrakk}R{\isacharsemicolon}\ Q{\isacharsemicolon}\ {\isasymnot}\ P{\isasymrbrakk}\ {\isasymLongrightarrow}\ R
   180 *}
   204 *}
   181 
   205 
   182   apply assumption
       
   183  apply assumption
       
   184 done
       
   185 
       
   186 
       
   187 
   206 
   188 text{*Quantifiers*}
   207 text{*Quantifiers*}
   189 
   208 
   190 lemma "\<forall>x. P x \<longrightarrow> P x"
   209 lemma "\<forall>x. P x \<longrightarrow> P x"
   191 apply (rule allI)
   210 apply (rule allI)
   192 apply (rule impI)
   211 by (rule impI)
   193 apply assumption
   212 text{*NEW*}
   194 done
       
   195 
   213 
   196 lemma "(\<forall>x. P \<longrightarrow> Q x) \<Longrightarrow> P \<longrightarrow> (\<forall>x. Q x)"
   214 lemma "(\<forall>x. P \<longrightarrow> Q x) \<Longrightarrow> P \<longrightarrow> (\<forall>x. Q x)"
   197 apply (rule impI)
   215 apply (rule impI, rule allI)
   198 apply (rule allI)
       
   199 apply (drule spec)
   216 apply (drule spec)
   200 apply (drule mp)
   217 by (drule mp)
   201   apply assumption
   218 text{*NEW*}
   202  apply assumption
   219 
   203 done
   220 lemma "\<lbrakk>\<forall>x. P x \<longrightarrow> P (h x); P a\<rbrakk> \<Longrightarrow> P(h (h a))"
   204 
       
   205 lemma "\<lbrakk>\<forall>x. P x \<longrightarrow> P (f x); P a\<rbrakk> \<Longrightarrow> P(f (f a))"
       
   206 apply (frule spec)
   221 apply (frule spec)
   207 apply (drule mp, assumption)
   222 apply (drule mp, assumption)
   208 apply (drule spec)
   223 apply (drule spec)
   209 apply (drule mp, assumption, assumption)
   224 by (drule mp)
   210 done
   225 text{*NEW*}
       
   226 
   211 
   227 
   212 text
   228 text
   213 {*
   229 {*
   214 proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ {\isadigit{1}}\isanewline
   230 proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ {\isadigit{1}}\isanewline
   215 \isanewline
   231 \isanewline
   226 \ {\isadigit{1}}{\isachardot}\ {\isasymlbrakk}{\isasymforall}x{\isachardot}\ P\ x\ {\isasymlongrightarrow}\ P\ {\isacharparenleft}f\ x{\isacharparenright}{\isacharsemicolon}\ P\ a{\isacharsemicolon}\ P\ {\isacharparenleft}f\ a{\isacharparenright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ {\isacharparenleft}f\ {\isacharparenleft}f\ a{\isacharparenright}{\isacharparenright}
   242 \ {\isadigit{1}}{\isachardot}\ {\isasymlbrakk}{\isasymforall}x{\isachardot}\ P\ x\ {\isasymlongrightarrow}\ P\ {\isacharparenleft}f\ x{\isacharparenright}{\isacharsemicolon}\ P\ a{\isacharsemicolon}\ P\ {\isacharparenleft}f\ a{\isacharparenright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ {\isacharparenleft}f\ {\isacharparenleft}f\ a{\isacharparenright}{\isacharparenright}
   227 *}
   243 *}
   228 
   244 
   229 lemma "\<lbrakk>\<forall>x. P x \<longrightarrow> P (f x); P a\<rbrakk> \<Longrightarrow> P(f (f a))"
   245 lemma "\<lbrakk>\<forall>x. P x \<longrightarrow> P (f x); P a\<rbrakk> \<Longrightarrow> P(f (f a))"
   230 by blast
   246 by blast
       
   247 
       
   248 text{*NEW
       
   249 Hilbert-epsilon theorems*}
       
   250 
       
   251 text{*
       
   252 @{thm[display] some_equality[no_vars]}
       
   253 \rulename{some_equality}
       
   254 
       
   255 @{thm[display] someI[no_vars]}
       
   256 \rulename{someI}
       
   257 
       
   258 @{thm[display] someI2[no_vars]}
       
   259 \rulename{someI2}
       
   260 
       
   261 needed for examples
       
   262 
       
   263 @{thm[display] inv_def[no_vars]}
       
   264 \rulename{inv_def}
       
   265 
       
   266 @{thm[display] Least_def[no_vars]}
       
   267 \rulename{Least_def}
       
   268 
       
   269 @{thm[display] order_antisym[no_vars]}
       
   270 \rulename{order_antisym}
       
   271 *}
       
   272 
       
   273 
       
   274 lemma "inv Suc (Suc x) = x"
       
   275 by (simp add: inv_def)
       
   276 
       
   277 text{*but we know nothing about inv Suc 0*}
       
   278 
       
   279 theorem Least_equality:
       
   280      "\<lbrakk> P (k::nat);  \<forall>x. P x \<longrightarrow> k \<le> x \<rbrakk> \<Longrightarrow> (LEAST x. P(x)) = k"
       
   281 apply (simp add: Least_def)
       
   282  
       
   283 txt{*omit maybe?
       
   284 @{subgoals[display,indent=0,margin=65]}
       
   285 *};
       
   286    
       
   287 apply (rule some_equality)
       
   288 
       
   289 txt{*
       
   290 @{subgoals[display,indent=0,margin=65]}
       
   291 
       
   292 first subgoal is existence; second is uniqueness
       
   293 *};
       
   294 by (auto intro: order_antisym)
       
   295 
       
   296 
       
   297 theorem axiom_of_choice:
       
   298      "(\<forall>x. \<exists> y. P x y) \<Longrightarrow> \<exists>f. \<forall>x. P x (f x)"
       
   299 apply (rule exI, rule allI)
       
   300 
       
   301 txt{*
       
   302 @{subgoals[display,indent=0,margin=65]}
       
   303 
       
   304 state after intro rules
       
   305 *};
       
   306 apply (drule spec, erule exE)
       
   307 
       
   308 txt{*
       
   309 @{subgoals[display,indent=0,margin=65]}
       
   310 
       
   311 applying @text{someI} automatically instantiates
       
   312 @{term f} to @{term "\<lambda>x. SOME y. P x y"}
       
   313 *};
       
   314 
       
   315 by (rule someI)
       
   316 
       
   317 (*both can be done by blast, which however hasn't been introduced yet*)
       
   318 lemma "[| P (k::nat);  \<forall>x. P x \<longrightarrow> k \<le> x |] ==> (LEAST x. P(x)) = k";
       
   319 apply (simp add: Least_def)
       
   320 by (blast intro: some_equality order_antisym);
       
   321 
       
   322 theorem axiom_of_choice: "(\<forall>x. \<exists> y. P x y) \<Longrightarrow> \<exists>f. \<forall>x. P x (f x)"
       
   323 apply (rule exI [of _  "\<lambda>x. SOME y. P x y"])
       
   324 by (blast intro: someI);
       
   325 
       
   326 text{*end of NEW material*}
   231 
   327 
   232 lemma "(\<exists>x. P x) \<or> (\<exists>x. Q x) \<Longrightarrow> \<exists>x. P x \<or> Q x"
   328 lemma "(\<exists>x. P x) \<or> (\<exists>x. Q x) \<Longrightarrow> \<exists>x. P x \<or> Q x"
   233 apply elim
   329 apply elim
   234  apply intro
   330  apply intro
   235  apply assumption
   331  apply assumption