doc-src/TutorialI/Rules/Basic.thy
changeset 11456 7eb63f63e6c6
parent 11407 138919f1a135
child 11458 09a6c44a48ea
equal deleted inserted replaced
11455:e07927b980ec 11456:7eb63f63e6c6
   307 
   307 
   308 text{*but we know nothing about inv Suc 0*}
   308 text{*but we know nothing about inv Suc 0*}
   309 
   309 
   310 theorem Least_equality:
   310 theorem Least_equality:
   311      "\<lbrakk> P (k::nat);  \<forall>x. P x \<longrightarrow> k \<le> x \<rbrakk> \<Longrightarrow> (LEAST x. P(x)) = k"
   311      "\<lbrakk> P (k::nat);  \<forall>x. P x \<longrightarrow> k \<le> x \<rbrakk> \<Longrightarrow> (LEAST x. P(x)) = k"
   312 apply (simp add: Least_def LeastM_def)
   312 apply (simp add: Least_def)
   313  
   313  
   314 txt{*omit maybe?
   314 txt{*omit maybe?
   315 @{subgoals[display,indent=0,margin=65]}
   315 @{subgoals[display,indent=0,margin=65]}
   316 *};
   316 *};
   317    
   317    
   318 apply (rule some_equality)
   318 apply (rule the_equality)
   319 
   319 
   320 txt{*
   320 txt{*
   321 @{subgoals[display,indent=0,margin=65]}
   321 @{subgoals[display,indent=0,margin=65]}
   322 
   322 
   323 first subgoal is existence; second is uniqueness
   323 first subgoal is existence; second is uniqueness