src/HOL/Nat.thy
changeset 26748 4d51ddd6aa5c
parent 26335 961bbcc9d85b
child 27104 791607529f6d
equal deleted inserted replaced
26747:f32fa5f5bdd1 26748:4d51ddd6aa5c
   732 
   732 
   733 lemma nat_mult_1_right: "n * (1::nat) = n"
   733 lemma nat_mult_1_right: "n * (1::nat) = n"
   734 by simp
   734 by simp
   735 
   735 
   736 
   736 
   737 subsubsection {* Additional theorems about "less than" *}
   737 subsubsection {* Additional theorems about @{term "op \<le>"} *}
       
   738 
       
   739 text {* Complete induction, aka course-of-values induction *}
       
   740 
       
   741 lemma less_induct [case_names less]:
       
   742   fixes P :: "nat \<Rightarrow> bool"
       
   743   assumes step: "\<And>x. (\<And>y. y < x \<Longrightarrow> P y) \<Longrightarrow> P x"
       
   744   shows "P a"
       
   745 proof - 
       
   746   have "\<And>z. z\<le>a \<Longrightarrow> P z"
       
   747   proof (induct a)
       
   748     case (0 z)
       
   749     have "P 0" by (rule step) auto
       
   750     thus ?case using 0 by auto
       
   751   next
       
   752     case (Suc x z)
       
   753     then have "z \<le> x \<or> z = Suc x" by (simp add: le_Suc_eq)
       
   754     thus ?case
       
   755     proof
       
   756       assume "z \<le> x" thus "P z" by (rule Suc(1))
       
   757     next
       
   758       assume z: "z = Suc x"
       
   759       show "P z"
       
   760         by (rule step) (rule Suc(1), simp add: z le_simps)
       
   761     qed
       
   762   qed
       
   763   thus ?thesis by auto
       
   764 qed
       
   765 
       
   766 lemma nat_less_induct:
       
   767   assumes "!!n. \<forall>m::nat. m < n --> P m ==> P n" shows "P n"
       
   768   using assms less_induct by blast
       
   769 
       
   770 lemma measure_induct_rule [case_names less]:
       
   771   fixes f :: "'a \<Rightarrow> nat"
       
   772   assumes step: "\<And>x. (\<And>y. f y < f x \<Longrightarrow> P y) \<Longrightarrow> P x"
       
   773   shows "P a"
       
   774 by (induct m\<equiv>"f a" arbitrary: a rule: less_induct) (auto intro: step)
       
   775 
       
   776 text {* old style induction rules: *}
       
   777 lemma measure_induct:
       
   778   fixes f :: "'a \<Rightarrow> nat"
       
   779   shows "(\<And>x. \<forall>y. f y < f x \<longrightarrow> P y \<Longrightarrow> P x) \<Longrightarrow> P a"
       
   780   by (rule measure_induct_rule [of f P a]) iprover
       
   781 
       
   782 lemma full_nat_induct:
       
   783   assumes step: "(!!n. (ALL m. Suc m <= n --> P m) ==> P n)"
       
   784   shows "P n"
       
   785   by (rule less_induct) (auto intro: step simp:le_simps)
   738 
   786 
   739 text{*An induction rule for estabilishing binary relations*}
   787 text{*An induction rule for estabilishing binary relations*}
   740 lemma less_Suc_induct:
   788 lemma less_Suc_induct:
   741   assumes less:  "i < j"
   789   assumes less:  "i < j"
   742      and  step:  "!!i. P i (Suc i)"
   790      and  step:  "!!i. P i (Suc i)"
   753     thus ?case by (auto intro: assms)
   801     thus ?case by (auto intro: assms)
   754   qed
   802   qed
   755   thus "P i j" by (simp add: j)
   803   thus "P i j" by (simp add: j)
   756 qed
   804 qed
   757 
   805 
       
   806 lemma nat_induct2: "[|P 0; P (Suc 0); !!k. P k ==> P (Suc (Suc k))|] ==> P n"
       
   807   apply (rule nat_less_induct)
       
   808   apply (case_tac n)
       
   809   apply (case_tac [2] nat)
       
   810   apply (blast intro: less_trans)+
       
   811   done
       
   812 
       
   813 text {* The method of infinite descent, frequently used in number theory.
       
   814 Provided by Roelof Oosterhuis.
       
   815 $P(n)$ is true for all $n\in\mathbb{N}$ if
       
   816 \begin{itemize}
       
   817   \item case ``0'': given $n=0$ prove $P(n)$,
       
   818   \item case ``smaller'': given $n>0$ and $\neg P(n)$ prove there exists
       
   819         a smaller integer $m$ such that $\neg P(m)$.
       
   820 \end{itemize} *}
       
   821 
       
   822 text{* A compact version without explicit base case: *}
       
   823 lemma infinite_descent:
       
   824   "\<lbrakk> !!n::nat. \<not> P n \<Longrightarrow>  \<exists>m<n. \<not>  P m \<rbrakk> \<Longrightarrow>  P n"
       
   825 by (induct n rule: less_induct, auto)
       
   826 
       
   827 lemma infinite_descent0[case_names 0 smaller]: 
       
   828   "\<lbrakk> P 0; !!n. n>0 \<Longrightarrow> \<not> P n \<Longrightarrow> (\<exists>m::nat. m < n \<and> \<not>P m) \<rbrakk> \<Longrightarrow> P n"
       
   829 by (rule infinite_descent) (case_tac "n>0", auto)
       
   830 
       
   831 text {*
       
   832 Infinite descent using a mapping to $\mathbb{N}$:
       
   833 $P(x)$ is true for all $x\in D$ if there exists a $V: D \to \mathbb{N}$ and
       
   834 \begin{itemize}
       
   835 \item case ``0'': given $V(x)=0$ prove $P(x)$,
       
   836 \item case ``smaller'': given $V(x)>0$ and $\neg P(x)$ prove there exists a $y \in D$ such that $V(y)<V(x)$ and $~\neg P(y)$.
       
   837 \end{itemize}
       
   838 NB: the proof also shows how to use the previous lemma. *}
       
   839 
       
   840 corollary infinite_descent0_measure [case_names 0 smaller]:
       
   841   assumes A0: "!!x. V x = (0::nat) \<Longrightarrow> P x"
       
   842     and   A1: "!!x. V x > 0 \<Longrightarrow> \<not>P x \<Longrightarrow> (\<exists>y. V y < V x \<and> \<not>P y)"
       
   843   shows "P x"
       
   844 proof -
       
   845   obtain n where "n = V x" by auto
       
   846   moreover have "\<And>x. V x = n \<Longrightarrow> P x"
       
   847   proof (induct n rule: infinite_descent0)
       
   848     case 0 -- "i.e. $V(x) = 0$"
       
   849     with A0 show "P x" by auto
       
   850   next -- "now $n>0$ and $P(x)$ does not hold for some $x$ with $V(x)=n$"
       
   851     case (smaller n)
       
   852     then obtain x where vxn: "V x = n " and "V x > 0 \<and> \<not> P x" by auto
       
   853     with A1 obtain y where "V y < V x \<and> \<not> P y" by auto
       
   854     with vxn obtain m where "m = V y \<and> m<n \<and> \<not> P y" by auto
       
   855     then show ?case by auto
       
   856   qed
       
   857   ultimately show "P x" by auto
       
   858 qed
       
   859 
       
   860 text{* Again, without explicit base case: *}
       
   861 lemma infinite_descent_measure:
       
   862 assumes "!!x. \<not> P x \<Longrightarrow> \<exists>y. (V::'a\<Rightarrow>nat) y < V x \<and> \<not> P y" shows "P x"
       
   863 proof -
       
   864   from assms obtain n where "n = V x" by auto
       
   865   moreover have "!!x. V x = n \<Longrightarrow> P x"
       
   866   proof (induct n rule: infinite_descent, auto)
       
   867     fix x assume "\<not> P x"
       
   868     with assms show "\<exists>m < V x. \<exists>y. V y = m \<and> \<not> P y" by auto
       
   869   qed
       
   870   ultimately show "P x" by auto
       
   871 qed
       
   872 
   758 text {* A [clumsy] way of lifting @{text "<"}
   873 text {* A [clumsy] way of lifting @{text "<"}
   759   monotonicity to @{text "\<le>"} monotonicity *}
   874   monotonicity to @{text "\<le>"} monotonicity *}
   760 lemma less_mono_imp_le_mono:
   875 lemma less_mono_imp_le_mono:
   761   "\<lbrakk> !!i j::nat. i < j \<Longrightarrow> f i < f j; i \<le> j \<rbrakk> \<Longrightarrow> f i \<le> ((f j)::nat)"
   876   "\<lbrakk> !!i j::nat. i < j \<Longrightarrow> f i < f j; i \<le> j \<rbrakk> \<Longrightarrow> f i \<le> ((f j)::nat)"
   762 by (simp add: order_le_less) (blast)
   877 by (simp add: order_le_less) (blast)
   807 apply (drule add_lessD1)
   922 apply (drule add_lessD1)
   808 apply (erule less_irrefl [THEN notE])
   923 apply (erule less_irrefl [THEN notE])
   809 done
   924 done
   810 
   925 
   811 lemma not_add_less2 [iff]: "~ (j + i < (i::nat))"
   926 lemma not_add_less2 [iff]: "~ (j + i < (i::nat))"
   812 by (simp add: add_commute not_add_less1)
   927 by (simp add: add_commute)
   813 
   928 
   814 lemma add_leD1: "m + k \<le> n ==> m \<le> (n::nat)"
   929 lemma add_leD1: "m + k \<le> n ==> m \<le> (n::nat)"
   815 apply (rule order_trans [of _ "m+k"])
   930 apply (rule order_trans [of _ "m+k"])
   816 apply (simp_all add: le_add1)
   931 apply (simp_all add: le_add1)
   817 done
   932 done
   839 
   954 
   840 lemma le_add_diff_inverse [simp]: "n \<le> m ==> n + (m - n) = (m::nat)"
   955 lemma le_add_diff_inverse [simp]: "n \<le> m ==> n + (m - n) = (m::nat)"
   841 by (simp add: add_diff_inverse linorder_not_less)
   956 by (simp add: add_diff_inverse linorder_not_less)
   842 
   957 
   843 lemma le_add_diff_inverse2 [simp]: "n \<le> m ==> (m - n) + n = (m::nat)"
   958 lemma le_add_diff_inverse2 [simp]: "n \<le> m ==> (m - n) + n = (m::nat)"
   844 by (simp add: le_add_diff_inverse add_commute)
   959 by (simp add: add_commute)
   845 
   960 
   846 lemma Suc_diff_le: "n \<le> m ==> Suc m - n = Suc (m - n)"
   961 lemma Suc_diff_le: "n \<le> m ==> Suc m - n = Suc (m - n)"
   847 by (induct m n rule: diff_induct) simp_all
   962 by (induct m n rule: diff_induct) simp_all
   848 
   963 
   849 lemma diff_less_Suc: "m - n < Suc m"
   964 lemma diff_less_Suc: "m - n < Suc m"
  1326 Least_Suc}, since there appears to be no need.*}
  1441 Least_Suc}, since there appears to be no need.*}
  1327 
  1442 
  1328 subsection {* size of a datatype value *}
  1443 subsection {* size of a datatype value *}
  1329 
  1444 
  1330 class size = type +
  1445 class size = type +
  1331   fixes size :: "'a \<Rightarrow> nat" -- {* see further theory @{text Wellfounded_Recursion} *}
  1446   fixes size :: "'a \<Rightarrow> nat" -- {* see further theory @{text Wellfounded} *}
  1332 
  1447 
  1333 end
  1448 end