--- a/doc-src/Locales/Locales/Examples3.thy Sun Jul 05 21:10:23 2009 +0200
+++ b/doc-src/Locales/Locales/Examples3.thy Thu Jul 09 08:55:42 2009 +0200
@@ -152,7 +152,7 @@
apply (rule_tac x = "gcd x y" in exI)
apply auto [1]
apply (rule_tac x = "lcm x y" in exI)
- apply (auto intro: nat_lcm_least)
+ apply (auto intro: lcm_least_nat)
done
then interpret nat_dvd: lattice "op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool" .
show "partial_order.less op dvd (x::nat) y = (x dvd y \<and> x \<noteq> y)"
@@ -168,7 +168,7 @@
apply (unfold nat_dvd.join_def)
apply (rule the_equality)
apply (unfold nat_dvd.is_sup_def)
- apply (auto intro: nat_lcm_least iff: nat_lcm_unique)
+ apply (auto intro: lcm_least_nat iff: lcm_unique_nat)
done
qed
--- a/doc-src/Locales/Locales/document/Examples3.tex Sun Jul 05 21:10:23 2009 +0200
+++ b/doc-src/Locales/Locales/document/Examples3.tex Thu Jul 09 08:55:42 2009 +0200
@@ -330,7 +330,7 @@
\ \ \ \ \isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}rule{\isacharunderscore}tac\ x\ {\isacharequal}\ {\isachardoublequoteopen}lcm\ x\ y{\isachardoublequoteclose}\ \isakeyword{in}\ exI{\isacharparenright}\isanewline
\ \ \ \ \isacommand{apply}\isamarkupfalse%
-\ {\isacharparenleft}auto\ intro{\isacharcolon}\ nat{\isacharunderscore}lcm{\isacharunderscore}least{\isacharparenright}\isanewline
+\ {\isacharparenleft}auto\ intro{\isacharcolon}\ lcm{\isacharunderscore}least{\isacharunderscore}nat{\isacharparenright}\isanewline
\ \ \ \ \isacommand{done}\isamarkupfalse%
\isanewline
\ \ \isacommand{then}\isamarkupfalse%
@@ -364,7 +364,7 @@
\ \ \ \ \isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}unfold\ nat{\isacharunderscore}dvd{\isachardot}is{\isacharunderscore}sup{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \isacommand{apply}\isamarkupfalse%
-\ {\isacharparenleft}auto\ intro{\isacharcolon}\ nat{\isacharunderscore}lcm{\isacharunderscore}least\ iff{\isacharcolon}\ nat{\isacharunderscore}lcm{\isacharunderscore}unique{\isacharparenright}\isanewline
+\ {\isacharparenleft}auto\ intro{\isacharcolon}\ lcm{\isacharunderscore}least{\isacharunderscore}nat\ iff{\isacharcolon}\ lcm{\isacharunderscore}unique{\isacharunderscore}nat{\isacharparenright}\isanewline
\ \ \ \ \isacommand{done}\isamarkupfalse%
\isanewline
\isacommand{qed}\isamarkupfalse%
--- a/src/HOL/Algebra/Exponent.thy Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/Algebra/Exponent.thy Thu Jul 09 08:55:42 2009 +0200
@@ -217,7 +217,7 @@
prefer 2 apply (blast intro: dvd_mult2)
apply (drule dvd_diffD1)
apply assumption
- prefer 2 apply (blast intro: nat_dvd_diff)
+ prefer 2 apply (blast intro: dvd_diff_nat)
apply (drule gr0_implies_Suc, auto)
done
@@ -233,7 +233,7 @@
prefer 2 apply (blast intro: dvd_mult2)
apply (drule dvd_diffD1)
apply assumption
- prefer 2 apply (blast intro: nat_dvd_diff)
+ prefer 2 apply (blast intro: dvd_diff_nat)
apply (drule less_imp_Suc_add, auto)
done
--- a/src/HOL/Bali/AxCompl.thy Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/Bali/AxCompl.thy Thu Jul 09 08:55:42 2009 +0200
@@ -1402,7 +1402,7 @@
apply -
apply (induct_tac "n")
apply (tactic "ALLGOALS (clarsimp_tac @{clasimpset})")
-apply (tactic {* dtac (permute_prems 0 1 (thm "card_seteq")) 1 *})
+apply (tactic {* dtac (Thm.permute_prems 0 1 (thm "card_seteq")) 1 *})
apply simp
apply (erule finite_imageI)
apply (simp add: MGF_asm ax_derivs_asm)
--- a/src/HOL/Decision_Procs/Cooper.thy Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/Decision_Procs/Cooper.thy Thu Jul 09 08:55:42 2009 +0200
@@ -1042,7 +1042,7 @@
consts
plusinf:: "fm \<Rightarrow> fm" (* Virtual substitution of +\<infinity>*)
minusinf:: "fm \<Rightarrow> fm" (* Virtual substitution of -\<infinity>*)
- \<delta> :: "fm \<Rightarrow> int" (* Compute lcm {d| N\<^isup>?\<^isup> Dvd c*x+t \<in> p}*)
+ \<delta> :: "fm \<Rightarrow> int" (* Compute lcm {d| N\<^isup>? Dvd c*x+t \<in> p}*)
d\<delta> :: "fm \<Rightarrow> int \<Rightarrow> bool" (* checks if a given l divides all the ds above*)
recdef minusinf "measure size"
@@ -1104,7 +1104,7 @@
proof (induct p rule: iszlfm.induct)
case (1 p q)
let ?d = "\<delta> (And p q)"
- from prems int_lcm_pos have dp: "?d >0" by simp
+ from prems lcm_pos_int have dp: "?d >0" by simp
have d1: "\<delta> p dvd \<delta> (And p q)" using prems by simp
hence th: "d\<delta> p ?d" using delta_mono prems(3-4) by(simp only: iszlfm.simps)
have "\<delta> q dvd \<delta> (And p q)" using prems by simp
@@ -1113,7 +1113,7 @@
next
case (2 p q)
let ?d = "\<delta> (And p q)"
- from prems int_lcm_pos have dp: "?d >0" by simp
+ from prems lcm_pos_int have dp: "?d >0" by simp
have "\<delta> p dvd \<delta> (And p q)" using prems by simp
hence th: "d\<delta> p ?d" using delta_mono prems by(simp only: iszlfm.simps)
have "\<delta> q dvd \<delta> (And p q)" using prems by simp
@@ -1410,15 +1410,15 @@
from prems have dl2: "\<zeta> q dvd lcm (\<zeta> p) (\<zeta> q)" by simp
from prems d\<beta>_mono[where p = "p" and l="\<zeta> p" and l'="lcm (\<zeta> p) (\<zeta> q)"]
d\<beta>_mono[where p = "q" and l="\<zeta> q" and l'="lcm (\<zeta> p) (\<zeta> q)"]
- dl1 dl2 show ?case by (auto simp add: int_lcm_pos)
+ dl1 dl2 show ?case by (auto simp add: lcm_pos_int)
next
case (2 p q)
from prems have dl1: "\<zeta> p dvd lcm (\<zeta> p) (\<zeta> q)" by simp
from prems have dl2: "\<zeta> q dvd lcm (\<zeta> p) (\<zeta> q)" by simp
from prems d\<beta>_mono[where p = "p" and l="\<zeta> p" and l'="lcm (\<zeta> p) (\<zeta> q)"]
d\<beta>_mono[where p = "q" and l="\<zeta> q" and l'="lcm (\<zeta> p) (\<zeta> q)"]
- dl1 dl2 show ?case by (auto simp add: int_lcm_pos)
-qed (auto simp add: int_lcm_pos)
+ dl1 dl2 show ?case by (auto simp add: lcm_pos_int)
+qed (auto simp add: lcm_pos_int)
lemma a\<beta>: assumes linp: "iszlfm p" and d: "d\<beta> p l" and lp: "l > 0"
shows "iszlfm (a\<beta> p l) \<and> d\<beta> (a\<beta> p l) 1 \<and> (Ifm bbs (l*x #bs) (a\<beta> p l) = Ifm bbs (x#bs) p)"
--- a/src/HOL/Decision_Procs/Ferrack.thy Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/Decision_Procs/Ferrack.thy Thu Jul 09 08:55:42 2009 +0200
@@ -746,7 +746,7 @@
moreover
{assume g1:"?g>1" hence g0: "?g > 0" by simp
from g1 nnz have gp0: "?g' \<noteq> 0" by simp
- hence g'p: "?g' > 0" using int_gcd_ge_0[where x="n" and y="numgcd ?t'"] by arith
+ hence g'p: "?g' > 0" using gcd_ge_0_int[where x="n" and y="numgcd ?t'"] by arith
hence "?g'= 1 \<or> ?g' > 1" by arith
moreover {assume "?g'=1" hence ?thesis by (simp add: Let_def simp_num_pair_def simpnum_ci)}
moreover {assume g'1:"?g'>1"
@@ -782,7 +782,7 @@
moreover
{assume g1:"?g>1" hence g0: "?g > 0" by simp
from g1 nnz have gp0: "?g' \<noteq> 0" by simp
- hence g'p: "?g' > 0" using int_gcd_ge_0[where x="n" and y="numgcd ?t'"] by arith
+ hence g'p: "?g' > 0" using gcd_ge_0_int[where x="n" and y="numgcd ?t'"] by arith
hence "?g'= 1 \<or> ?g' > 1" by arith
moreover {assume "?g'=1" hence ?thesis using prems
by (auto simp add: Let_def simp_num_pair_def simpnum_numbound0)}
--- a/src/HOL/Decision_Procs/MIR.thy Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/Decision_Procs/MIR.thy Thu Jul 09 08:55:42 2009 +0200
@@ -1060,7 +1060,7 @@
moreover
{assume g1:"?g>1" hence g0: "?g > 0" by simp
from g1 nnz have gp0: "?g' \<noteq> 0" by simp
- hence g'p: "?g' > 0" using int_gcd_ge_0[where x="n" and y="numgcd ?t'"] by arith
+ hence g'p: "?g' > 0" using gcd_ge_0_int[where x="n" and y="numgcd ?t'"] by arith
hence "?g'= 1 \<or> ?g' > 1" by arith
moreover {assume "?g'=1" hence ?thesis by (simp add: Let_def simp_num_pair_def)}
moreover {assume g'1:"?g'>1"
@@ -1096,7 +1096,7 @@
moreover
{assume g1:"?g>1" hence g0: "?g > 0" by simp
from g1 nnz have gp0: "?g' \<noteq> 0" by simp
- hence g'p: "?g' > 0" using int_gcd_ge_0[where x="n" and y="numgcd ?t'"] by arith
+ hence g'p: "?g' > 0" using gcd_ge_0_int[where x="n" and y="numgcd ?t'"] by arith
hence "?g'= 1 \<or> ?g' > 1" by arith
moreover {assume "?g'=1" hence ?thesis using prems
by (auto simp add: Let_def simp_num_pair_def)}
@@ -1233,7 +1233,7 @@
moreover
{assume g1:"?g>1" hence g0: "?g > 0" by simp
from g1 dnz have gp0: "?g' \<noteq> 0" by simp
- hence g'p: "?g' > 0" using int_gcd_ge_0[where x="d" and y="numgcd t"] by arith
+ hence g'p: "?g' > 0" using gcd_ge_0_int[where x="d" and y="numgcd t"] by arith
hence "?g'= 1 \<or> ?g' > 1" by arith
moreover {assume "?g'=1" hence ?thesis by (simp add: Let_def simpdvd_def)}
moreover {assume g'1:"?g'>1"
@@ -2126,7 +2126,7 @@
proof (induct p rule: iszlfm.induct)
case (1 p q)
let ?d = "\<delta> (And p q)"
- from prems int_lcm_pos have dp: "?d >0" by simp
+ from prems lcm_pos_int have dp: "?d >0" by simp
have d1: "\<delta> p dvd \<delta> (And p q)" using prems by simp
hence th: "d\<delta> p ?d"
using delta_mono prems by(simp only: iszlfm.simps) blast
@@ -2136,7 +2136,7 @@
next
case (2 p q)
let ?d = "\<delta> (And p q)"
- from prems int_lcm_pos have dp: "?d >0" by simp
+ from prems lcm_pos_int have dp: "?d >0" by simp
have "\<delta> p dvd \<delta> (And p q)" using prems by simp hence th: "d\<delta> p ?d" using delta_mono prems
by(simp only: iszlfm.simps) blast
have "\<delta> q dvd \<delta> (And p q)" using prems by simp hence th': "d\<delta> q ?d" using delta_mono prems by(simp only: iszlfm.simps) blast
@@ -2514,15 +2514,15 @@
from prems have dl2: "\<zeta> q dvd lcm (\<zeta> p) (\<zeta> q)" by simp
from prems d\<beta>_mono[where p = "p" and l="\<zeta> p" and l'="lcm (\<zeta> p) (\<zeta> q)"]
d\<beta>_mono[where p = "q" and l="\<zeta> q" and l'="lcm (\<zeta> p) (\<zeta> q)"]
- dl1 dl2 show ?case by (auto simp add: int_lcm_pos)
+ dl1 dl2 show ?case by (auto simp add: lcm_pos_int)
next
case (2 p q)
from prems have dl1: "\<zeta> p dvd lcm (\<zeta> p) (\<zeta> q)" by simp
from prems have dl2: "\<zeta> q dvd lcm (\<zeta> p) (\<zeta> q)" by simp
from prems d\<beta>_mono[where p = "p" and l="\<zeta> p" and l'="lcm (\<zeta> p) (\<zeta> q)"]
d\<beta>_mono[where p = "q" and l="\<zeta> q" and l'="lcm (\<zeta> p) (\<zeta> q)"]
- dl1 dl2 show ?case by (auto simp add: int_lcm_pos)
-qed (auto simp add: int_lcm_pos)
+ dl1 dl2 show ?case by (auto simp add: lcm_pos_int)
+qed (auto simp add: lcm_pos_int)
lemma a\<beta>: assumes linp: "iszlfm p (a #bs)" and d: "d\<beta> p l" and lp: "l > 0"
shows "iszlfm (a\<beta> p l) (a #bs) \<and> d\<beta> (a\<beta> p l) 1 \<and> (Ifm (real (l * x) #bs) (a\<beta> p l) = Ifm ((real x)#bs) p)"
@@ -4175,7 +4175,7 @@
by (induct p rule: isrlfm.induct, auto)
lemma zgcd_le1: assumes ip: "(0::int) < i" shows "gcd i j \<le> i"
proof-
- from int_gcd_dvd1 have th: "gcd i j dvd i" by blast
+ from gcd_dvd1_int have th: "gcd i j dvd i" by blast
from zdvd_imp_le[OF th ip] show ?thesis .
qed
--- a/src/HOL/Divides.thy Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/Divides.thy Thu Jul 09 08:55:42 2009 +0200
@@ -887,7 +887,7 @@
interpretation dvd: order "op dvd" "\<lambda>n m \<Colon> nat. n dvd m \<and> \<not> m dvd n"
proof qed (auto intro: dvd_refl dvd_trans dvd_anti_sym)
-lemma nat_dvd_diff[simp]: "[| k dvd m; k dvd n |] ==> k dvd (m-n :: nat)"
+lemma dvd_diff_nat[simp]: "[| k dvd m; k dvd n |] ==> k dvd (m-n :: nat)"
unfolding dvd_def
by (blast intro: diff_mult_distrib2 [symmetric])
@@ -897,7 +897,7 @@
done
lemma dvd_diffD1: "[| k dvd m-n; k dvd m; n\<le>m |] ==> k dvd (n::nat)"
-by (drule_tac m = m in nat_dvd_diff, auto)
+by (drule_tac m = m in dvd_diff_nat, auto)
lemma dvd_reduce: "(k dvd n + k) = (k dvd (n::nat))"
apply (rule iffI)
@@ -906,7 +906,7 @@
apply (subgoal_tac "n = (n+k) -k")
prefer 2 apply simp
apply (erule ssubst)
- apply (erule nat_dvd_diff)
+ apply (erule dvd_diff_nat)
apply (rule dvd_refl)
done
--- a/src/HOL/Extraction/Euclid.thy Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/Extraction/Euclid.thy Thu Jul 09 08:55:42 2009 +0200
@@ -189,7 +189,7 @@
assume pn: "p \<le> n"
from `prime p` have "0 < p" by (rule prime_g_zero)
then have "p dvd n!" using pn by (rule dvd_factorial)
- with dvd have "p dvd ?k - n!" by (rule nat_dvd_diff)
+ with dvd have "p dvd ?k - n!" by (rule dvd_diff_nat)
then have "p dvd 1" by simp
with prime show False using prime_nd_one by auto
qed
--- a/src/HOL/Fun.thy Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/Fun.thy Thu Jul 09 08:55:42 2009 +0200
@@ -496,6 +496,40 @@
hide (open) const swap
+
+subsection {* Inversion of injective functions *}
+
+definition inv :: "('a \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'a)" where
+ "inv f y = (THE x. f x = y)"
+
+lemma inv_f_f:
+ assumes "inj f"
+ shows "inv f (f x) = x"
+proof -
+ from assms have "(THE x'. f x' = f x) = (THE x'. x' = x)"
+ by (simp only: inj_eq)
+ also have "... = x" by (rule the_eq_trivial)
+ finally show ?thesis by (unfold inv_def)
+qed
+
+lemma f_inv_f:
+ assumes "inj f"
+ and "y \<in> range f"
+ shows "f (inv f y) = y"
+proof (unfold inv_def)
+ from `y \<in> range f` obtain x where "y = f x" ..
+ then have "f x = y" ..
+ then show "f (THE x. f x = y) = y"
+ proof (rule theI)
+ fix x' assume "f x' = y"
+ with `f x = y` have "f x' = f x" by simp
+ with `inj f` show "x' = x" by (rule injD)
+ qed
+qed
+
+hide (open) const inv
+
+
subsection {* Proof tool setup *}
text {* simplifies terms of the form
--- a/src/HOL/GCD.thy Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/GCD.thy Thu Jul 09 08:55:42 2009 +0200
@@ -163,7 +163,7 @@
subsection {* GCD *}
(* was gcd_induct *)
-lemma nat_gcd_induct:
+lemma gcd_nat_induct:
fixes m n :: nat
assumes "\<And>m. P m 0"
and "\<And>m n. 0 < n \<Longrightarrow> P n (m mod n) \<Longrightarrow> P m n"
@@ -175,43 +175,43 @@
(* specific to int *)
-lemma int_gcd_neg1 [simp]: "gcd (-x::int) y = gcd x y"
+lemma gcd_neg1_int [simp]: "gcd (-x::int) y = gcd x y"
by (simp add: gcd_int_def)
-lemma int_gcd_neg2 [simp]: "gcd (x::int) (-y) = gcd x y"
+lemma gcd_neg2_int [simp]: "gcd (x::int) (-y) = gcd x y"
by (simp add: gcd_int_def)
lemma abs_gcd_int[simp]: "abs(gcd (x::int) y) = gcd x y"
by(simp add: gcd_int_def)
-lemma int_gcd_abs: "gcd (x::int) y = gcd (abs x) (abs y)"
+lemma gcd_abs_int: "gcd (x::int) y = gcd (abs x) (abs y)"
by (simp add: gcd_int_def)
lemma gcd_abs1_int[simp]: "gcd (abs x) (y::int) = gcd x y"
-by (metis abs_idempotent int_gcd_abs)
+by (metis abs_idempotent gcd_abs_int)
lemma gcd_abs2_int[simp]: "gcd x (abs y::int) = gcd x y"
-by (metis abs_idempotent int_gcd_abs)
+by (metis abs_idempotent gcd_abs_int)
-lemma int_gcd_cases:
+lemma gcd_cases_int:
fixes x :: int and y
assumes "x >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> P (gcd x y)"
and "x >= 0 \<Longrightarrow> y <= 0 \<Longrightarrow> P (gcd x (-y))"
and "x <= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> P (gcd (-x) y)"
and "x <= 0 \<Longrightarrow> y <= 0 \<Longrightarrow> P (gcd (-x) (-y))"
shows "P (gcd x y)"
-by (insert prems, auto simp add: int_gcd_neg1 int_gcd_neg2, arith)
+by (insert prems, auto simp add: gcd_neg1_int gcd_neg2_int, arith)
-lemma int_gcd_ge_0 [simp]: "gcd (x::int) y >= 0"
+lemma gcd_ge_0_int [simp]: "gcd (x::int) y >= 0"
by (simp add: gcd_int_def)
-lemma int_lcm_neg1: "lcm (-x::int) y = lcm x y"
+lemma lcm_neg1_int: "lcm (-x::int) y = lcm x y"
by (simp add: lcm_int_def)
-lemma int_lcm_neg2: "lcm (x::int) (-y) = lcm x y"
+lemma lcm_neg2_int: "lcm (x::int) (-y) = lcm x y"
by (simp add: lcm_int_def)
-lemma int_lcm_abs: "lcm (x::int) y = lcm (abs x) (abs y)"
+lemma lcm_abs_int: "lcm (x::int) y = lcm (abs x) (abs y)"
by (simp add: lcm_int_def)
lemma abs_lcm_int [simp]: "abs (lcm i j::int) = lcm i j"
@@ -223,53 +223,53 @@
lemma lcm_abs2_int[simp]: "lcm x (abs y::int) = lcm x y"
by (metis abs_idempotent lcm_int_def)
-lemma int_lcm_cases:
+lemma lcm_cases_int:
fixes x :: int and y
assumes "x >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> P (lcm x y)"
and "x >= 0 \<Longrightarrow> y <= 0 \<Longrightarrow> P (lcm x (-y))"
and "x <= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> P (lcm (-x) y)"
and "x <= 0 \<Longrightarrow> y <= 0 \<Longrightarrow> P (lcm (-x) (-y))"
shows "P (lcm x y)"
-by (insert prems, auto simp add: int_lcm_neg1 int_lcm_neg2, arith)
+by (insert prems, auto simp add: lcm_neg1_int lcm_neg2_int, arith)
-lemma int_lcm_ge_0 [simp]: "lcm (x::int) y >= 0"
+lemma lcm_ge_0_int [simp]: "lcm (x::int) y >= 0"
by (simp add: lcm_int_def)
(* was gcd_0, etc. *)
-lemma nat_gcd_0 [simp]: "gcd (x::nat) 0 = x"
+lemma gcd_0_nat [simp]: "gcd (x::nat) 0 = x"
by simp
(* was igcd_0, etc. *)
-lemma int_gcd_0 [simp]: "gcd (x::int) 0 = abs x"
+lemma gcd_0_int [simp]: "gcd (x::int) 0 = abs x"
by (unfold gcd_int_def, auto)
-lemma nat_gcd_0_left [simp]: "gcd 0 (x::nat) = x"
+lemma gcd_0_left_nat [simp]: "gcd 0 (x::nat) = x"
by simp
-lemma int_gcd_0_left [simp]: "gcd 0 (x::int) = abs x"
+lemma gcd_0_left_int [simp]: "gcd 0 (x::int) = abs x"
by (unfold gcd_int_def, auto)
-lemma nat_gcd_red: "gcd (x::nat) y = gcd y (x mod y)"
+lemma gcd_red_nat: "gcd (x::nat) y = gcd y (x mod y)"
by (case_tac "y = 0", auto)
(* weaker, but useful for the simplifier *)
-lemma nat_gcd_non_0: "y ~= (0::nat) \<Longrightarrow> gcd (x::nat) y = gcd y (x mod y)"
+lemma gcd_non_0_nat: "y ~= (0::nat) \<Longrightarrow> gcd (x::nat) y = gcd y (x mod y)"
by simp
-lemma nat_gcd_1 [simp]: "gcd (m::nat) 1 = 1"
+lemma gcd_1_nat [simp]: "gcd (m::nat) 1 = 1"
by simp
-lemma nat_gcd_Suc_0 [simp]: "gcd (m::nat) (Suc 0) = Suc 0"
+lemma gcd_Suc_0 [simp]: "gcd (m::nat) (Suc 0) = Suc 0"
by (simp add: One_nat_def)
-lemma int_gcd_1 [simp]: "gcd (m::int) 1 = 1"
+lemma gcd_1_int [simp]: "gcd (m::int) 1 = 1"
by (simp add: gcd_int_def)
-lemma nat_gcd_idem: "gcd (x::nat) x = x"
+lemma gcd_idem_nat: "gcd (x::nat) x = x"
by simp
-lemma int_gcd_idem: "gcd (x::int) x = abs x"
+lemma gcd_idem_int: "gcd (x::int) x = abs x"
by (auto simp add: gcd_int_def)
declare gcd_nat.simps [simp del]
@@ -279,260 +279,260 @@
conjunctions don't seem provable separately.
*}
-lemma nat_gcd_dvd1 [iff]: "(gcd (m::nat)) n dvd m"
- and nat_gcd_dvd2 [iff]: "(gcd m n) dvd n"
- apply (induct m n rule: nat_gcd_induct)
- apply (simp_all add: nat_gcd_non_0)
+lemma gcd_dvd1_nat [iff]: "(gcd (m::nat)) n dvd m"
+ and gcd_dvd2_nat [iff]: "(gcd m n) dvd n"
+ apply (induct m n rule: gcd_nat_induct)
+ apply (simp_all add: gcd_non_0_nat)
apply (blast dest: dvd_mod_imp_dvd)
done
-lemma int_gcd_dvd1 [iff]: "gcd (x::int) y dvd x"
-by (metis gcd_int_def int_dvd_iff nat_gcd_dvd1)
+lemma gcd_dvd1_int [iff]: "gcd (x::int) y dvd x"
+by (metis gcd_int_def int_dvd_iff gcd_dvd1_nat)
-lemma int_gcd_dvd2 [iff]: "gcd (x::int) y dvd y"
-by (metis gcd_int_def int_dvd_iff nat_gcd_dvd2)
+lemma gcd_dvd2_int [iff]: "gcd (x::int) y dvd y"
+by (metis gcd_int_def int_dvd_iff gcd_dvd2_nat)
lemma dvd_gcd_D1_nat: "k dvd gcd m n \<Longrightarrow> (k::nat) dvd m"
-by(metis nat_gcd_dvd1 dvd_trans)
+by(metis gcd_dvd1_nat dvd_trans)
lemma dvd_gcd_D2_nat: "k dvd gcd m n \<Longrightarrow> (k::nat) dvd n"
-by(metis nat_gcd_dvd2 dvd_trans)
+by(metis gcd_dvd2_nat dvd_trans)
lemma dvd_gcd_D1_int: "i dvd gcd m n \<Longrightarrow> (i::int) dvd m"
-by(metis int_gcd_dvd1 dvd_trans)
+by(metis gcd_dvd1_int dvd_trans)
lemma dvd_gcd_D2_int: "i dvd gcd m n \<Longrightarrow> (i::int) dvd n"
-by(metis int_gcd_dvd2 dvd_trans)
+by(metis gcd_dvd2_int dvd_trans)
-lemma nat_gcd_le1 [simp]: "a \<noteq> 0 \<Longrightarrow> gcd (a::nat) b \<le> a"
+lemma gcd_le1_nat [simp]: "a \<noteq> 0 \<Longrightarrow> gcd (a::nat) b \<le> a"
by (rule dvd_imp_le, auto)
-lemma nat_gcd_le2 [simp]: "b \<noteq> 0 \<Longrightarrow> gcd (a::nat) b \<le> b"
+lemma gcd_le2_nat [simp]: "b \<noteq> 0 \<Longrightarrow> gcd (a::nat) b \<le> b"
by (rule dvd_imp_le, auto)
-lemma int_gcd_le1 [simp]: "a > 0 \<Longrightarrow> gcd (a::int) b \<le> a"
+lemma gcd_le1_int [simp]: "a > 0 \<Longrightarrow> gcd (a::int) b \<le> a"
by (rule zdvd_imp_le, auto)
-lemma int_gcd_le2 [simp]: "b > 0 \<Longrightarrow> gcd (a::int) b \<le> b"
+lemma gcd_le2_int [simp]: "b > 0 \<Longrightarrow> gcd (a::int) b \<le> b"
by (rule zdvd_imp_le, auto)
-lemma nat_gcd_greatest: "(k::nat) dvd m \<Longrightarrow> k dvd n \<Longrightarrow> k dvd gcd m n"
-by (induct m n rule: nat_gcd_induct) (simp_all add: nat_gcd_non_0 dvd_mod)
+lemma gcd_greatest_nat: "(k::nat) dvd m \<Longrightarrow> k dvd n \<Longrightarrow> k dvd gcd m n"
+by (induct m n rule: gcd_nat_induct) (simp_all add: gcd_non_0_nat dvd_mod)
-lemma int_gcd_greatest:
+lemma gcd_greatest_int:
"(k::int) dvd m \<Longrightarrow> k dvd n \<Longrightarrow> k dvd gcd m n"
- apply (subst int_gcd_abs)
+ apply (subst gcd_abs_int)
apply (subst abs_dvd_iff [symmetric])
- apply (rule nat_gcd_greatest [transferred])
+ apply (rule gcd_greatest_nat [transferred])
apply auto
done
-lemma nat_gcd_greatest_iff [iff]: "(k dvd gcd (m::nat) n) =
+lemma gcd_greatest_iff_nat [iff]: "(k dvd gcd (m::nat) n) =
(k dvd m & k dvd n)"
- by (blast intro!: nat_gcd_greatest intro: dvd_trans)
+ by (blast intro!: gcd_greatest_nat intro: dvd_trans)
-lemma int_gcd_greatest_iff: "((k::int) dvd gcd m n) = (k dvd m & k dvd n)"
- by (blast intro!: int_gcd_greatest intro: dvd_trans)
+lemma gcd_greatest_iff_int: "((k::int) dvd gcd m n) = (k dvd m & k dvd n)"
+ by (blast intro!: gcd_greatest_int intro: dvd_trans)
-lemma nat_gcd_zero [simp]: "(gcd (m::nat) n = 0) = (m = 0 & n = 0)"
- by (simp only: dvd_0_left_iff [symmetric] nat_gcd_greatest_iff)
+lemma gcd_zero_nat [simp]: "(gcd (m::nat) n = 0) = (m = 0 & n = 0)"
+ by (simp only: dvd_0_left_iff [symmetric] gcd_greatest_iff_nat)
-lemma int_gcd_zero [simp]: "(gcd (m::int) n = 0) = (m = 0 & n = 0)"
+lemma gcd_zero_int [simp]: "(gcd (m::int) n = 0) = (m = 0 & n = 0)"
by (auto simp add: gcd_int_def)
-lemma nat_gcd_pos [simp]: "(gcd (m::nat) n > 0) = (m ~= 0 | n ~= 0)"
- by (insert nat_gcd_zero [of m n], arith)
+lemma gcd_pos_nat [simp]: "(gcd (m::nat) n > 0) = (m ~= 0 | n ~= 0)"
+ by (insert gcd_zero_nat [of m n], arith)
-lemma int_gcd_pos [simp]: "(gcd (m::int) n > 0) = (m ~= 0 | n ~= 0)"
- by (insert int_gcd_zero [of m n], insert int_gcd_ge_0 [of m n], arith)
+lemma gcd_pos_int [simp]: "(gcd (m::int) n > 0) = (m ~= 0 | n ~= 0)"
+ by (insert gcd_zero_int [of m n], insert gcd_ge_0_int [of m n], arith)
-lemma nat_gcd_commute: "gcd (m::nat) n = gcd n m"
+lemma gcd_commute_nat: "gcd (m::nat) n = gcd n m"
by (rule dvd_anti_sym, auto)
-lemma int_gcd_commute: "gcd (m::int) n = gcd n m"
- by (auto simp add: gcd_int_def nat_gcd_commute)
+lemma gcd_commute_int: "gcd (m::int) n = gcd n m"
+ by (auto simp add: gcd_int_def gcd_commute_nat)
-lemma nat_gcd_assoc: "gcd (gcd (k::nat) m) n = gcd k (gcd m n)"
+lemma gcd_assoc_nat: "gcd (gcd (k::nat) m) n = gcd k (gcd m n)"
apply (rule dvd_anti_sym)
apply (blast intro: dvd_trans)+
done
-lemma int_gcd_assoc: "gcd (gcd (k::int) m) n = gcd k (gcd m n)"
- by (auto simp add: gcd_int_def nat_gcd_assoc)
+lemma gcd_assoc_int: "gcd (gcd (k::int) m) n = gcd k (gcd m n)"
+ by (auto simp add: gcd_int_def gcd_assoc_nat)
-lemmas nat_gcd_left_commute =
- mk_left_commute[of gcd, OF nat_gcd_assoc nat_gcd_commute]
+lemmas gcd_left_commute_nat =
+ mk_left_commute[of gcd, OF gcd_assoc_nat gcd_commute_nat]
-lemmas int_gcd_left_commute =
- mk_left_commute[of gcd, OF int_gcd_assoc int_gcd_commute]
+lemmas gcd_left_commute_int =
+ mk_left_commute[of gcd, OF gcd_assoc_int gcd_commute_int]
-lemmas nat_gcd_ac = nat_gcd_assoc nat_gcd_commute nat_gcd_left_commute
+lemmas gcd_ac_nat = gcd_assoc_nat gcd_commute_nat gcd_left_commute_nat
-- {* gcd is an AC-operator *}
-lemmas int_gcd_ac = int_gcd_assoc int_gcd_commute int_gcd_left_commute
+lemmas gcd_ac_int = gcd_assoc_int gcd_commute_int gcd_left_commute_int
-lemma nat_gcd_unique: "(d::nat) dvd a \<and> d dvd b \<and>
+lemma gcd_unique_nat: "(d::nat) dvd a \<and> d dvd b \<and>
(\<forall>e. e dvd a \<and> e dvd b \<longrightarrow> e dvd d) \<longleftrightarrow> d = gcd a b"
apply auto
apply (rule dvd_anti_sym)
- apply (erule (1) nat_gcd_greatest)
+ apply (erule (1) gcd_greatest_nat)
apply auto
done
-lemma int_gcd_unique: "d >= 0 & (d::int) dvd a \<and> d dvd b \<and>
+lemma gcd_unique_int: "d >= 0 & (d::int) dvd a \<and> d dvd b \<and>
(\<forall>e. e dvd a \<and> e dvd b \<longrightarrow> e dvd d) \<longleftrightarrow> d = gcd a b"
apply (case_tac "d = 0")
apply force
apply (rule iffI)
apply (rule zdvd_anti_sym)
apply arith
- apply (subst int_gcd_pos)
+ apply (subst gcd_pos_int)
apply clarsimp
apply (drule_tac x = "d + 1" in spec)
apply (frule zdvd_imp_le)
- apply (auto intro: int_gcd_greatest)
+ apply (auto intro: gcd_greatest_int)
done
lemma gcd_proj1_if_dvd_nat [simp]: "(x::nat) dvd y \<Longrightarrow> gcd x y = x"
-by (metis dvd.eq_iff nat_gcd_unique)
+by (metis dvd.eq_iff gcd_unique_nat)
lemma gcd_proj2_if_dvd_nat [simp]: "(y::nat) dvd x \<Longrightarrow> gcd x y = y"
-by (metis dvd.eq_iff nat_gcd_unique)
+by (metis dvd.eq_iff gcd_unique_nat)
lemma gcd_proj1_if_dvd_int[simp]: "x dvd y \<Longrightarrow> gcd (x::int) y = abs x"
-by (metis abs_dvd_iff abs_eq_0 int_gcd_0_left int_gcd_abs int_gcd_unique)
+by (metis abs_dvd_iff abs_eq_0 gcd_0_left_int gcd_abs_int gcd_unique_int)
lemma gcd_proj2_if_dvd_int[simp]: "y dvd x \<Longrightarrow> gcd (x::int) y = abs y"
-by (metis gcd_proj1_if_dvd_int int_gcd_commute)
+by (metis gcd_proj1_if_dvd_int gcd_commute_int)
text {*
\medskip Multiplication laws
*}
-lemma nat_gcd_mult_distrib: "(k::nat) * gcd m n = gcd (k * m) (k * n)"
+lemma gcd_mult_distrib_nat: "(k::nat) * gcd m n = gcd (k * m) (k * n)"
-- {* \cite[page 27]{davenport92} *}
- apply (induct m n rule: nat_gcd_induct)
+ apply (induct m n rule: gcd_nat_induct)
apply simp
apply (case_tac "k = 0")
- apply (simp_all add: mod_geq nat_gcd_non_0 mod_mult_distrib2)
+ apply (simp_all add: mod_geq gcd_non_0_nat mod_mult_distrib2)
done
-lemma int_gcd_mult_distrib: "abs (k::int) * gcd m n = gcd (k * m) (k * n)"
- apply (subst (1 2) int_gcd_abs)
+lemma gcd_mult_distrib_int: "abs (k::int) * gcd m n = gcd (k * m) (k * n)"
+ apply (subst (1 2) gcd_abs_int)
apply (subst (1 2) abs_mult)
- apply (rule nat_gcd_mult_distrib [transferred])
+ apply (rule gcd_mult_distrib_nat [transferred])
apply auto
done
-lemma nat_coprime_dvd_mult: "coprime (k::nat) n \<Longrightarrow> k dvd m * n \<Longrightarrow> k dvd m"
- apply (insert nat_gcd_mult_distrib [of m k n])
+lemma coprime_dvd_mult_nat: "coprime (k::nat) n \<Longrightarrow> k dvd m * n \<Longrightarrow> k dvd m"
+ apply (insert gcd_mult_distrib_nat [of m k n])
apply simp
apply (erule_tac t = m in ssubst)
apply simp
done
-lemma int_coprime_dvd_mult:
+lemma coprime_dvd_mult_int:
"coprime (k::int) n \<Longrightarrow> k dvd m * n \<Longrightarrow> k dvd m"
apply (subst abs_dvd_iff [symmetric])
apply (subst dvd_abs_iff [symmetric])
-apply (subst (asm) int_gcd_abs)
-apply (rule nat_coprime_dvd_mult [transferred])
+apply (subst (asm) gcd_abs_int)
+apply (rule coprime_dvd_mult_nat [transferred])
prefer 4 apply assumption
apply auto
apply (subst abs_mult [symmetric], auto)
done
-lemma nat_coprime_dvd_mult_iff: "coprime (k::nat) n \<Longrightarrow>
+lemma coprime_dvd_mult_iff_nat: "coprime (k::nat) n \<Longrightarrow>
(k dvd m * n) = (k dvd m)"
- by (auto intro: nat_coprime_dvd_mult)
+ by (auto intro: coprime_dvd_mult_nat)
-lemma int_coprime_dvd_mult_iff: "coprime (k::int) n \<Longrightarrow>
+lemma coprime_dvd_mult_iff_int: "coprime (k::int) n \<Longrightarrow>
(k dvd m * n) = (k dvd m)"
- by (auto intro: int_coprime_dvd_mult)
+ by (auto intro: coprime_dvd_mult_int)
-lemma nat_gcd_mult_cancel: "coprime k n \<Longrightarrow> gcd ((k::nat) * m) n = gcd m n"
+lemma gcd_mult_cancel_nat: "coprime k n \<Longrightarrow> gcd ((k::nat) * m) n = gcd m n"
apply (rule dvd_anti_sym)
- apply (rule nat_gcd_greatest)
- apply (rule_tac n = k in nat_coprime_dvd_mult)
- apply (simp add: nat_gcd_assoc)
- apply (simp add: nat_gcd_commute)
+ apply (rule gcd_greatest_nat)
+ apply (rule_tac n = k in coprime_dvd_mult_nat)
+ apply (simp add: gcd_assoc_nat)
+ apply (simp add: gcd_commute_nat)
apply (simp_all add: mult_commute)
done
-lemma int_gcd_mult_cancel:
+lemma gcd_mult_cancel_int:
"coprime (k::int) n \<Longrightarrow> gcd (k * m) n = gcd m n"
-apply (subst (1 2) int_gcd_abs)
+apply (subst (1 2) gcd_abs_int)
apply (subst abs_mult)
-apply (rule nat_gcd_mult_cancel [transferred], auto)
+apply (rule gcd_mult_cancel_nat [transferred], auto)
done
text {* \medskip Addition laws *}
-lemma nat_gcd_add1 [simp]: "gcd ((m::nat) + n) n = gcd m n"
+lemma gcd_add1_nat [simp]: "gcd ((m::nat) + n) n = gcd m n"
apply (case_tac "n = 0")
- apply (simp_all add: nat_gcd_non_0)
+ apply (simp_all add: gcd_non_0_nat)
done
-lemma nat_gcd_add2 [simp]: "gcd (m::nat) (m + n) = gcd m n"
- apply (subst (1 2) nat_gcd_commute)
+lemma gcd_add2_nat [simp]: "gcd (m::nat) (m + n) = gcd m n"
+ apply (subst (1 2) gcd_commute_nat)
apply (subst add_commute)
apply simp
done
(* to do: add the other variations? *)
-lemma nat_gcd_diff1: "(m::nat) >= n \<Longrightarrow> gcd (m - n) n = gcd m n"
- by (subst nat_gcd_add1 [symmetric], auto)
+lemma gcd_diff1_nat: "(m::nat) >= n \<Longrightarrow> gcd (m - n) n = gcd m n"
+ by (subst gcd_add1_nat [symmetric], auto)
-lemma nat_gcd_diff2: "(n::nat) >= m \<Longrightarrow> gcd (n - m) n = gcd m n"
- apply (subst nat_gcd_commute)
- apply (subst nat_gcd_diff1 [symmetric])
+lemma gcd_diff2_nat: "(n::nat) >= m \<Longrightarrow> gcd (n - m) n = gcd m n"
+ apply (subst gcd_commute_nat)
+ apply (subst gcd_diff1_nat [symmetric])
apply auto
- apply (subst nat_gcd_commute)
- apply (subst nat_gcd_diff1)
+ apply (subst gcd_commute_nat)
+ apply (subst gcd_diff1_nat)
apply assumption
- apply (rule nat_gcd_commute)
+ apply (rule gcd_commute_nat)
done
-lemma int_gcd_non_0: "(y::int) > 0 \<Longrightarrow> gcd x y = gcd y (x mod y)"
+lemma gcd_non_0_int: "(y::int) > 0 \<Longrightarrow> gcd x y = gcd y (x mod y)"
apply (frule_tac b = y and a = x in pos_mod_sign)
apply (simp del: pos_mod_sign add: gcd_int_def abs_if nat_mod_distrib)
- apply (auto simp add: nat_gcd_non_0 nat_mod_distrib [symmetric]
+ apply (auto simp add: gcd_non_0_nat nat_mod_distrib [symmetric]
zmod_zminus1_eq_if)
apply (frule_tac a = x in pos_mod_bound)
- apply (subst (1 2) nat_gcd_commute)
- apply (simp del: pos_mod_bound add: nat_diff_distrib nat_gcd_diff2
+ apply (subst (1 2) gcd_commute_nat)
+ apply (simp del: pos_mod_bound add: nat_diff_distrib gcd_diff2_nat
nat_le_eq_zle)
done
-lemma int_gcd_red: "gcd (x::int) y = gcd y (x mod y)"
+lemma gcd_red_int: "gcd (x::int) y = gcd y (x mod y)"
apply (case_tac "y = 0")
apply force
apply (case_tac "y > 0")
- apply (subst int_gcd_non_0, auto)
- apply (insert int_gcd_non_0 [of "-y" "-x"])
- apply (auto simp add: int_gcd_neg1 int_gcd_neg2)
+ apply (subst gcd_non_0_int, auto)
+ apply (insert gcd_non_0_int [of "-y" "-x"])
+ apply (auto simp add: gcd_neg1_int gcd_neg2_int)
done
-lemma int_gcd_add1 [simp]: "gcd ((m::int) + n) n = gcd m n"
-by (metis int_gcd_red mod_add_self1 zadd_commute)
+lemma gcd_add1_int [simp]: "gcd ((m::int) + n) n = gcd m n"
+by (metis gcd_red_int mod_add_self1 zadd_commute)
-lemma int_gcd_add2 [simp]: "gcd m ((m::int) + n) = gcd m n"
-by (metis int_gcd_add1 int_gcd_commute zadd_commute)
+lemma gcd_add2_int [simp]: "gcd m ((m::int) + n) = gcd m n"
+by (metis gcd_add1_int gcd_commute_int zadd_commute)
-lemma nat_gcd_add_mult: "gcd (m::nat) (k * m + n) = gcd m n"
-by (metis mod_mult_self3 nat_gcd_commute nat_gcd_red)
+lemma gcd_add_mult_nat: "gcd (m::nat) (k * m + n) = gcd m n"
+by (metis mod_mult_self3 gcd_commute_nat gcd_red_nat)
-lemma int_gcd_add_mult: "gcd (m::int) (k * m + n) = gcd m n"
-by (metis int_gcd_commute int_gcd_red mod_mult_self1 zadd_commute)
+lemma gcd_add_mult_int: "gcd (m::int) (k * m + n) = gcd m n"
+by (metis gcd_commute_int gcd_red_int mod_mult_self1 zadd_commute)
(* to do: differences, and all variations of addition rules
as simplification rules for nat and int *)
(* FIXME remove iff *)
-lemma nat_gcd_dvd_prod [iff]: "gcd (m::nat) n dvd k * n"
+lemma gcd_dvd_prod_nat [iff]: "gcd (m::nat) n dvd k * n"
using mult_dvd_mono [of 1] by auto
(* to do: add the three variations of these, and for ints? *)
@@ -559,7 +559,7 @@
apply(rule Max_eqI[THEN sym])
apply (metis dvd.eq_iff finite_Collect_conjI finite_divisors_nat)
apply simp
- apply(metis Suc_diff_1 Suc_neq_Zero dvd_imp_le nat_gcd_greatest_iff nat_gcd_pos)
+ apply(metis Suc_diff_1 Suc_neq_Zero dvd_imp_le gcd_greatest_iff_nat gcd_pos_nat)
apply simp
done
@@ -568,14 +568,14 @@
apply(rule Max_eqI[THEN sym])
apply (metis dvd.eq_iff finite_Collect_conjI finite_divisors_int)
apply simp
- apply (metis int_gcd_greatest_iff int_gcd_pos zdvd_imp_le)
+ apply (metis gcd_greatest_iff_int gcd_pos_int zdvd_imp_le)
apply simp
done
subsection {* Coprimality *}
-lemma nat_div_gcd_coprime:
+lemma div_gcd_coprime_nat:
assumes nz: "(a::nat) \<noteq> 0 \<or> b \<noteq> 0"
shows "coprime (a div gcd a b) (b div gcd a b)"
proof -
@@ -593,34 +593,34 @@
then have dvdgg':"?g * ?g' dvd a" "?g* ?g' dvd b"
by (auto simp add: dvd_mult_div_cancel [OF dvdg(1)]
dvd_mult_div_cancel [OF dvdg(2)] dvd_def)
- have "?g \<noteq> 0" using nz by (simp add: nat_gcd_zero)
+ have "?g \<noteq> 0" using nz by (simp add: gcd_zero_nat)
then have gp: "?g > 0" by arith
- from nat_gcd_greatest [OF dvdgg'] have "?g * ?g' dvd ?g" .
+ from gcd_greatest_nat [OF dvdgg'] have "?g * ?g' dvd ?g" .
with dvd_mult_cancel1 [OF gp] show "?g' = 1" by simp
qed
-lemma int_div_gcd_coprime:
+lemma div_gcd_coprime_int:
assumes nz: "(a::int) \<noteq> 0 \<or> b \<noteq> 0"
shows "coprime (a div gcd a b) (b div gcd a b)"
-apply (subst (1 2 3) int_gcd_abs)
+apply (subst (1 2 3) gcd_abs_int)
apply (subst (1 2) abs_div)
apply simp
apply simp
apply(subst (1 2) abs_gcd_int)
-apply (rule nat_div_gcd_coprime [transferred])
-using nz apply (auto simp add: int_gcd_abs [symmetric])
+apply (rule div_gcd_coprime_nat [transferred])
+using nz apply (auto simp add: gcd_abs_int [symmetric])
done
-lemma nat_coprime: "coprime (a::nat) b \<longleftrightarrow> (\<forall>d. d dvd a \<and> d dvd b \<longleftrightarrow> d = 1)"
- using nat_gcd_unique[of 1 a b, simplified] by auto
+lemma coprime_nat: "coprime (a::nat) b \<longleftrightarrow> (\<forall>d. d dvd a \<and> d dvd b \<longleftrightarrow> d = 1)"
+ using gcd_unique_nat[of 1 a b, simplified] by auto
-lemma nat_coprime_Suc_0:
+lemma coprime_Suc_0_nat:
"coprime (a::nat) b \<longleftrightarrow> (\<forall>d. d dvd a \<and> d dvd b \<longleftrightarrow> d = Suc 0)"
- using nat_coprime by (simp add: One_nat_def)
+ using coprime_nat by (simp add: One_nat_def)
-lemma int_coprime: "coprime (a::int) b \<longleftrightarrow>
+lemma coprime_int: "coprime (a::int) b \<longleftrightarrow>
(\<forall>d. d >= 0 \<and> d dvd a \<and> d dvd b \<longleftrightarrow> d = 1)"
- using int_gcd_unique [of 1 a b]
+ using gcd_unique_int [of 1 a b]
apply clarsimp
apply (erule subst)
apply (rule iffI)
@@ -631,7 +631,7 @@
apply force
done
-lemma nat_gcd_coprime:
+lemma gcd_coprime_nat:
assumes z: "gcd (a::nat) b \<noteq> 0" and a: "a = a' * gcd a b" and
b: "b = b' * gcd a b"
shows "coprime a' b'"
@@ -640,7 +640,7 @@
apply (erule ssubst)
apply (subgoal_tac "b' = b div gcd a b")
apply (erule ssubst)
- apply (rule nat_div_gcd_coprime)
+ apply (rule div_gcd_coprime_nat)
using prems
apply force
apply (subst (1) b)
@@ -649,7 +649,7 @@
using z apply force
done
-lemma int_gcd_coprime:
+lemma gcd_coprime_int:
assumes z: "gcd (a::int) b \<noteq> 0" and a: "a = a' * gcd a b" and
b: "b = b' * gcd a b"
shows "coprime a' b'"
@@ -658,7 +658,7 @@
apply (erule ssubst)
apply (subgoal_tac "b' = b div gcd a b")
apply (erule ssubst)
- apply (rule int_div_gcd_coprime)
+ apply (rule div_gcd_coprime_int)
using prems
apply force
apply (subst (1) b)
@@ -667,117 +667,117 @@
using z apply force
done
-lemma nat_coprime_mult: assumes da: "coprime (d::nat) a" and db: "coprime d b"
+lemma coprime_mult_nat: assumes da: "coprime (d::nat) a" and db: "coprime d b"
shows "coprime d (a * b)"
- apply (subst nat_gcd_commute)
- using da apply (subst nat_gcd_mult_cancel)
- apply (subst nat_gcd_commute, assumption)
- apply (subst nat_gcd_commute, rule db)
+ apply (subst gcd_commute_nat)
+ using da apply (subst gcd_mult_cancel_nat)
+ apply (subst gcd_commute_nat, assumption)
+ apply (subst gcd_commute_nat, rule db)
done
-lemma int_coprime_mult: assumes da: "coprime (d::int) a" and db: "coprime d b"
+lemma coprime_mult_int: assumes da: "coprime (d::int) a" and db: "coprime d b"
shows "coprime d (a * b)"
- apply (subst int_gcd_commute)
- using da apply (subst int_gcd_mult_cancel)
- apply (subst int_gcd_commute, assumption)
- apply (subst int_gcd_commute, rule db)
+ apply (subst gcd_commute_int)
+ using da apply (subst gcd_mult_cancel_int)
+ apply (subst gcd_commute_int, assumption)
+ apply (subst gcd_commute_int, rule db)
done
-lemma nat_coprime_lmult:
+lemma coprime_lmult_nat:
assumes dab: "coprime (d::nat) (a * b)" shows "coprime d a"
proof -
have "gcd d a dvd gcd d (a * b)"
- by (rule nat_gcd_greatest, auto)
+ by (rule gcd_greatest_nat, auto)
with dab show ?thesis
by auto
qed
-lemma int_coprime_lmult:
+lemma coprime_lmult_int:
assumes "coprime (d::int) (a * b)" shows "coprime d a"
proof -
have "gcd d a dvd gcd d (a * b)"
- by (rule int_gcd_greatest, auto)
+ by (rule gcd_greatest_int, auto)
with assms show ?thesis
by auto
qed
-lemma nat_coprime_rmult:
+lemma coprime_rmult_nat:
assumes "coprime (d::nat) (a * b)" shows "coprime d b"
proof -
have "gcd d b dvd gcd d (a * b)"
- by (rule nat_gcd_greatest, auto intro: dvd_mult)
+ by (rule gcd_greatest_nat, auto intro: dvd_mult)
with assms show ?thesis
by auto
qed
-lemma int_coprime_rmult:
+lemma coprime_rmult_int:
assumes dab: "coprime (d::int) (a * b)" shows "coprime d b"
proof -
have "gcd d b dvd gcd d (a * b)"
- by (rule int_gcd_greatest, auto intro: dvd_mult)
+ by (rule gcd_greatest_int, auto intro: dvd_mult)
with dab show ?thesis
by auto
qed
-lemma nat_coprime_mul_eq: "coprime (d::nat) (a * b) \<longleftrightarrow>
+lemma coprime_mul_eq_nat: "coprime (d::nat) (a * b) \<longleftrightarrow>
coprime d a \<and> coprime d b"
- using nat_coprime_rmult[of d a b] nat_coprime_lmult[of d a b]
- nat_coprime_mult[of d a b]
+ using coprime_rmult_nat[of d a b] coprime_lmult_nat[of d a b]
+ coprime_mult_nat[of d a b]
by blast
-lemma int_coprime_mul_eq: "coprime (d::int) (a * b) \<longleftrightarrow>
+lemma coprime_mul_eq_int: "coprime (d::int) (a * b) \<longleftrightarrow>
coprime d a \<and> coprime d b"
- using int_coprime_rmult[of d a b] int_coprime_lmult[of d a b]
- int_coprime_mult[of d a b]
+ using coprime_rmult_int[of d a b] coprime_lmult_int[of d a b]
+ coprime_mult_int[of d a b]
by blast
-lemma nat_gcd_coprime_exists:
+lemma gcd_coprime_exists_nat:
assumes nz: "gcd (a::nat) b \<noteq> 0"
shows "\<exists>a' b'. a = a' * gcd a b \<and> b = b' * gcd a b \<and> coprime a' b'"
apply (rule_tac x = "a div gcd a b" in exI)
apply (rule_tac x = "b div gcd a b" in exI)
- using nz apply (auto simp add: nat_div_gcd_coprime dvd_div_mult)
+ using nz apply (auto simp add: div_gcd_coprime_nat dvd_div_mult)
done
-lemma int_gcd_coprime_exists:
+lemma gcd_coprime_exists_int:
assumes nz: "gcd (a::int) b \<noteq> 0"
shows "\<exists>a' b'. a = a' * gcd a b \<and> b = b' * gcd a b \<and> coprime a' b'"
apply (rule_tac x = "a div gcd a b" in exI)
apply (rule_tac x = "b div gcd a b" in exI)
- using nz apply (auto simp add: int_div_gcd_coprime dvd_div_mult_self)
+ using nz apply (auto simp add: div_gcd_coprime_int dvd_div_mult_self)
done
-lemma nat_coprime_exp: "coprime (d::nat) a \<Longrightarrow> coprime d (a^n)"
- by (induct n, simp_all add: nat_coprime_mult)
+lemma coprime_exp_nat: "coprime (d::nat) a \<Longrightarrow> coprime d (a^n)"
+ by (induct n, simp_all add: coprime_mult_nat)
-lemma int_coprime_exp: "coprime (d::int) a \<Longrightarrow> coprime d (a^n)"
- by (induct n, simp_all add: int_coprime_mult)
+lemma coprime_exp_int: "coprime (d::int) a \<Longrightarrow> coprime d (a^n)"
+ by (induct n, simp_all add: coprime_mult_int)
-lemma nat_coprime_exp2 [intro]: "coprime (a::nat) b \<Longrightarrow> coprime (a^n) (b^m)"
- apply (rule nat_coprime_exp)
- apply (subst nat_gcd_commute)
- apply (rule nat_coprime_exp)
- apply (subst nat_gcd_commute, assumption)
+lemma coprime_exp2_nat [intro]: "coprime (a::nat) b \<Longrightarrow> coprime (a^n) (b^m)"
+ apply (rule coprime_exp_nat)
+ apply (subst gcd_commute_nat)
+ apply (rule coprime_exp_nat)
+ apply (subst gcd_commute_nat, assumption)
done
-lemma int_coprime_exp2 [intro]: "coprime (a::int) b \<Longrightarrow> coprime (a^n) (b^m)"
- apply (rule int_coprime_exp)
- apply (subst int_gcd_commute)
- apply (rule int_coprime_exp)
- apply (subst int_gcd_commute, assumption)
+lemma coprime_exp2_int [intro]: "coprime (a::int) b \<Longrightarrow> coprime (a^n) (b^m)"
+ apply (rule coprime_exp_int)
+ apply (subst gcd_commute_int)
+ apply (rule coprime_exp_int)
+ apply (subst gcd_commute_int, assumption)
done
-lemma nat_gcd_exp: "gcd ((a::nat)^n) (b^n) = (gcd a b)^n"
+lemma gcd_exp_nat: "gcd ((a::nat)^n) (b^n) = (gcd a b)^n"
proof (cases)
assume "a = 0 & b = 0"
thus ?thesis by simp
next assume "~(a = 0 & b = 0)"
hence "coprime ((a div gcd a b)^n) ((b div gcd a b)^n)"
- by (auto simp:nat_div_gcd_coprime)
+ by (auto simp:div_gcd_coprime_nat)
hence "gcd ((a div gcd a b)^n * (gcd a b)^n)
((b div gcd a b)^n * (gcd a b)^n) = (gcd a b)^n"
apply (subst (1 2) mult_commute)
- apply (subst nat_gcd_mult_distrib [symmetric])
+ apply (subst gcd_mult_distrib_nat [symmetric])
apply simp
done
also have "(a div gcd a b)^n * (gcd a b)^n = a^n"
@@ -797,29 +797,29 @@
finally show ?thesis .
qed
-lemma int_gcd_exp: "gcd ((a::int)^n) (b^n) = (gcd a b)^n"
- apply (subst (1 2) int_gcd_abs)
+lemma gcd_exp_int: "gcd ((a::int)^n) (b^n) = (gcd a b)^n"
+ apply (subst (1 2) gcd_abs_int)
apply (subst (1 2) power_abs)
- apply (rule nat_gcd_exp [where n = n, transferred])
+ apply (rule gcd_exp_nat [where n = n, transferred])
apply auto
done
-lemma nat_coprime_divprod: "(d::nat) dvd a * b \<Longrightarrow> coprime d a \<Longrightarrow> d dvd b"
- using nat_coprime_dvd_mult_iff[of d a b]
+lemma coprime_divprod_nat: "(d::nat) dvd a * b \<Longrightarrow> coprime d a \<Longrightarrow> d dvd b"
+ using coprime_dvd_mult_iff_nat[of d a b]
by (auto simp add: mult_commute)
-lemma int_coprime_divprod: "(d::int) dvd a * b \<Longrightarrow> coprime d a \<Longrightarrow> d dvd b"
- using int_coprime_dvd_mult_iff[of d a b]
+lemma coprime_divprod_int: "(d::int) dvd a * b \<Longrightarrow> coprime d a \<Longrightarrow> d dvd b"
+ using coprime_dvd_mult_iff_int[of d a b]
by (auto simp add: mult_commute)
-lemma nat_division_decomp: assumes dc: "(a::nat) dvd b * c"
+lemma division_decomp_nat: assumes dc: "(a::nat) dvd b * c"
shows "\<exists>b' c'. a = b' * c' \<and> b' dvd b \<and> c' dvd c"
proof-
let ?g = "gcd a b"
{assume "?g = 0" with dc have ?thesis by auto}
moreover
{assume z: "?g \<noteq> 0"
- from nat_gcd_coprime_exists[OF z]
+ from gcd_coprime_exists_nat[OF z]
obtain a' b' where ab': "a = a' * ?g" "b = b' * ?g" "coprime a' b'"
by blast
have thb: "?g dvd b" by auto
@@ -828,21 +828,21 @@
from dc ab'(1,2) have "a'*?g dvd (b'*?g) *c" by auto
hence "?g*a' dvd ?g * (b' * c)" by (simp add: mult_assoc)
with z have th_1: "a' dvd b' * c" by auto
- from nat_coprime_dvd_mult[OF ab'(3)] th_1
+ from coprime_dvd_mult_nat[OF ab'(3)] th_1
have thc: "a' dvd c" by (subst (asm) mult_commute, blast)
from ab' have "a = ?g*a'" by algebra
with thb thc have ?thesis by blast }
ultimately show ?thesis by blast
qed
-lemma int_division_decomp: assumes dc: "(a::int) dvd b * c"
+lemma division_decomp_int: assumes dc: "(a::int) dvd b * c"
shows "\<exists>b' c'. a = b' * c' \<and> b' dvd b \<and> c' dvd c"
proof-
let ?g = "gcd a b"
{assume "?g = 0" with dc have ?thesis by auto}
moreover
{assume z: "?g \<noteq> 0"
- from int_gcd_coprime_exists[OF z]
+ from gcd_coprime_exists_int[OF z]
obtain a' b' where ab': "a = a' * ?g" "b = b' * ?g" "coprime a' b'"
by blast
have thb: "?g dvd b" by auto
@@ -852,14 +852,14 @@
from dc ab'(1,2) have "a'*?g dvd (b'*?g) *c" by auto
hence "?g*a' dvd ?g * (b' * c)" by (simp add: mult_assoc)
with z have th_1: "a' dvd b' * c" by auto
- from int_coprime_dvd_mult[OF ab'(3)] th_1
+ from coprime_dvd_mult_int[OF ab'(3)] th_1
have thc: "a' dvd c" by (subst (asm) mult_commute, blast)
from ab' have "a = ?g*a'" by algebra
with thb thc have ?thesis by blast }
ultimately show ?thesis by blast
qed
-lemma nat_pow_divides_pow:
+lemma pow_divides_pow_nat:
assumes ab: "(a::nat) ^ n dvd b ^n" and n:"n \<noteq> 0"
shows "a dvd b"
proof-
@@ -869,7 +869,7 @@
moreover
{assume z: "?g \<noteq> 0"
hence zn: "?g ^ n \<noteq> 0" using n by (simp add: neq0_conv)
- from nat_gcd_coprime_exists[OF z]
+ from gcd_coprime_exists_nat[OF z]
obtain a' b' where ab': "a = a' * ?g" "b = b' * ?g" "coprime a' b'"
by blast
from ab have "(a' * ?g) ^ n dvd (b' * ?g)^n"
@@ -880,14 +880,14 @@
have "a' dvd a'^n" by (simp add: m)
with th0 have "a' dvd b'^n" using dvd_trans[of a' "a'^n" "b'^n"] by simp
hence th1: "a' dvd b'^m * b'" by (simp add: m mult_commute)
- from nat_coprime_dvd_mult[OF nat_coprime_exp [OF ab'(3), of m]] th1
+ from coprime_dvd_mult_nat[OF coprime_exp_nat [OF ab'(3), of m]] th1
have "a' dvd b'" by (subst (asm) mult_commute, blast)
hence "a'*?g dvd b'*?g" by simp
with ab'(1,2) have ?thesis by simp }
ultimately show ?thesis by blast
qed
-lemma int_pow_divides_pow:
+lemma pow_divides_pow_int:
assumes ab: "(a::int) ^ n dvd b ^n" and n:"n \<noteq> 0"
shows "a dvd b"
proof-
@@ -897,7 +897,7 @@
moreover
{assume z: "?g \<noteq> 0"
hence zn: "?g ^ n \<noteq> 0" using n by (simp add: neq0_conv)
- from int_gcd_coprime_exists[OF z]
+ from gcd_coprime_exists_int[OF z]
obtain a' b' where ab': "a = a' * ?g" "b = b' * ?g" "coprime a' b'"
by blast
from ab have "(a' * ?g) ^ n dvd (b' * ?g)^n"
@@ -909,7 +909,7 @@
with th0 have "a' dvd b'^n"
using dvd_trans[of a' "a'^n" "b'^n"] by simp
hence th1: "a' dvd b'^m * b'" by (simp add: m mult_commute)
- from int_coprime_dvd_mult[OF int_coprime_exp [OF ab'(3), of m]] th1
+ from coprime_dvd_mult_int[OF coprime_exp_int [OF ab'(3), of m]] th1
have "a' dvd b'" by (subst (asm) mult_commute, blast)
hence "a'*?g dvd b'*?g" by simp
with ab'(1,2) have ?thesis by simp }
@@ -917,76 +917,76 @@
qed
(* FIXME move to Divides(?) *)
-lemma nat_pow_divides_eq [simp]: "n ~= 0 \<Longrightarrow> ((a::nat)^n dvd b^n) = (a dvd b)"
- by (auto intro: nat_pow_divides_pow dvd_power_same)
+lemma pow_divides_eq_nat [simp]: "n ~= 0 \<Longrightarrow> ((a::nat)^n dvd b^n) = (a dvd b)"
+ by (auto intro: pow_divides_pow_nat dvd_power_same)
-lemma int_pow_divides_eq [simp]: "n ~= 0 \<Longrightarrow> ((a::int)^n dvd b^n) = (a dvd b)"
- by (auto intro: int_pow_divides_pow dvd_power_same)
+lemma pow_divides_eq_int [simp]: "n ~= 0 \<Longrightarrow> ((a::int)^n dvd b^n) = (a dvd b)"
+ by (auto intro: pow_divides_pow_int dvd_power_same)
-lemma nat_divides_mult:
+lemma divides_mult_nat:
assumes mr: "(m::nat) dvd r" and nr: "n dvd r" and mn:"coprime m n"
shows "m * n dvd r"
proof-
from mr nr obtain m' n' where m': "r = m*m'" and n': "r = n*n'"
unfolding dvd_def by blast
from mr n' have "m dvd n'*n" by (simp add: mult_commute)
- hence "m dvd n'" using nat_coprime_dvd_mult_iff[OF mn] by simp
+ hence "m dvd n'" using coprime_dvd_mult_iff_nat[OF mn] by simp
then obtain k where k: "n' = m*k" unfolding dvd_def by blast
from n' k show ?thesis unfolding dvd_def by auto
qed
-lemma int_divides_mult:
+lemma divides_mult_int:
assumes mr: "(m::int) dvd r" and nr: "n dvd r" and mn:"coprime m n"
shows "m * n dvd r"
proof-
from mr nr obtain m' n' where m': "r = m*m'" and n': "r = n*n'"
unfolding dvd_def by blast
from mr n' have "m dvd n'*n" by (simp add: mult_commute)
- hence "m dvd n'" using int_coprime_dvd_mult_iff[OF mn] by simp
+ hence "m dvd n'" using coprime_dvd_mult_iff_int[OF mn] by simp
then obtain k where k: "n' = m*k" unfolding dvd_def by blast
from n' k show ?thesis unfolding dvd_def by auto
qed
-lemma nat_coprime_plus_one [simp]: "coprime ((n::nat) + 1) n"
+lemma coprime_plus_one_nat [simp]: "coprime ((n::nat) + 1) n"
apply (subgoal_tac "gcd (n + 1) n dvd (n + 1 - n)")
apply force
- apply (rule nat_dvd_diff)
+ apply (rule dvd_diff_nat)
apply auto
done
-lemma nat_coprime_Suc [simp]: "coprime (Suc n) n"
- using nat_coprime_plus_one by (simp add: One_nat_def)
+lemma coprime_Suc_nat [simp]: "coprime (Suc n) n"
+ using coprime_plus_one_nat by (simp add: One_nat_def)
-lemma int_coprime_plus_one [simp]: "coprime ((n::int) + 1) n"
+lemma coprime_plus_one_int [simp]: "coprime ((n::int) + 1) n"
apply (subgoal_tac "gcd (n + 1) n dvd (n + 1 - n)")
apply force
apply (rule dvd_diff)
apply auto
done
-lemma nat_coprime_minus_one: "(n::nat) \<noteq> 0 \<Longrightarrow> coprime (n - 1) n"
- using nat_coprime_plus_one [of "n - 1"]
- nat_gcd_commute [of "n - 1" n] by auto
+lemma coprime_minus_one_nat: "(n::nat) \<noteq> 0 \<Longrightarrow> coprime (n - 1) n"
+ using coprime_plus_one_nat [of "n - 1"]
+ gcd_commute_nat [of "n - 1" n] by auto
-lemma int_coprime_minus_one: "coprime ((n::int) - 1) n"
- using int_coprime_plus_one [of "n - 1"]
- int_gcd_commute [of "n - 1" n] by auto
+lemma coprime_minus_one_int: "coprime ((n::int) - 1) n"
+ using coprime_plus_one_int [of "n - 1"]
+ gcd_commute_int [of "n - 1" n] by auto
-lemma nat_setprod_coprime [rule_format]:
+lemma setprod_coprime_nat [rule_format]:
"(ALL i: A. coprime (f i) (x::nat)) --> coprime (PROD i:A. f i) x"
apply (case_tac "finite A")
apply (induct set: finite)
- apply (auto simp add: nat_gcd_mult_cancel)
+ apply (auto simp add: gcd_mult_cancel_nat)
done
-lemma int_setprod_coprime [rule_format]:
+lemma setprod_coprime_int [rule_format]:
"(ALL i: A. coprime (f i) (x::int)) --> coprime (PROD i:A. f i) x"
apply (case_tac "finite A")
apply (induct set: finite)
- apply (auto simp add: int_gcd_mult_cancel)
+ apply (auto simp add: gcd_mult_cancel_int)
done
-lemma nat_prime_odd: "prime (p::nat) \<Longrightarrow> p > 2 \<Longrightarrow> odd p"
+lemma prime_odd_nat: "prime (p::nat) \<Longrightarrow> p > 2 \<Longrightarrow> odd p"
unfolding prime_nat_def
apply (subst even_mult_two_ex)
apply clarify
@@ -994,41 +994,41 @@
apply auto
done
-lemma int_prime_odd: "prime (p::int) \<Longrightarrow> p > 2 \<Longrightarrow> odd p"
+lemma prime_odd_int: "prime (p::int) \<Longrightarrow> p > 2 \<Longrightarrow> odd p"
unfolding prime_int_def
- apply (frule nat_prime_odd)
+ apply (frule prime_odd_nat)
apply (auto simp add: even_nat_def)
done
-lemma nat_coprime_common_divisor: "coprime (a::nat) b \<Longrightarrow> x dvd a \<Longrightarrow>
+lemma coprime_common_divisor_nat: "coprime (a::nat) b \<Longrightarrow> x dvd a \<Longrightarrow>
x dvd b \<Longrightarrow> x = 1"
apply (subgoal_tac "x dvd gcd a b")
apply simp
- apply (erule (1) nat_gcd_greatest)
+ apply (erule (1) gcd_greatest_nat)
done
-lemma int_coprime_common_divisor: "coprime (a::int) b \<Longrightarrow> x dvd a \<Longrightarrow>
+lemma coprime_common_divisor_int: "coprime (a::int) b \<Longrightarrow> x dvd a \<Longrightarrow>
x dvd b \<Longrightarrow> abs x = 1"
apply (subgoal_tac "x dvd gcd a b")
apply simp
- apply (erule (1) int_gcd_greatest)
+ apply (erule (1) gcd_greatest_int)
done
-lemma nat_coprime_divisors: "(d::int) dvd a \<Longrightarrow> e dvd b \<Longrightarrow> coprime a b \<Longrightarrow>
+lemma coprime_divisors_nat: "(d::int) dvd a \<Longrightarrow> e dvd b \<Longrightarrow> coprime a b \<Longrightarrow>
coprime d e"
apply (auto simp add: dvd_def)
- apply (frule int_coprime_lmult)
- apply (subst int_gcd_commute)
- apply (subst (asm) (2) int_gcd_commute)
- apply (erule int_coprime_lmult)
+ apply (frule coprime_lmult_int)
+ apply (subst gcd_commute_int)
+ apply (subst (asm) (2) gcd_commute_int)
+ apply (erule coprime_lmult_int)
done
-lemma nat_invertible_coprime: "(x::nat) * y mod m = 1 \<Longrightarrow> coprime x m"
-apply (metis nat_coprime_lmult nat_gcd_1 nat_gcd_commute nat_gcd_red)
+lemma invertible_coprime_nat: "(x::nat) * y mod m = 1 \<Longrightarrow> coprime x m"
+apply (metis coprime_lmult_nat gcd_1_nat gcd_commute_nat gcd_red_nat)
done
-lemma int_invertible_coprime: "(x::int) * y mod m = 1 \<Longrightarrow> coprime x m"
-apply (metis int_coprime_lmult int_gcd_1 int_gcd_commute int_gcd_red)
+lemma invertible_coprime_int: "(x::int) * y mod m = 1 \<Longrightarrow> coprime x m"
+apply (metis coprime_lmult_int gcd_1_int gcd_commute_int gcd_red_int)
done
@@ -1054,7 +1054,7 @@
lemma bezw_aux [rule_format]:
"fst (bezw x y) * int x + snd (bezw x y) * int y = int (gcd x y)"
-proof (induct x y rule: nat_gcd_induct)
+proof (induct x y rule: gcd_nat_induct)
fix m :: nat
show "fst (bezw m 0) * int m + snd (bezw m 0) * int 0 = int (gcd m 0)"
by auto
@@ -1064,7 +1064,7 @@
snd (bezw n (m mod n)) * int (m mod n) =
int (gcd n (m mod n))"
thus "fst (bezw m n) * int m + snd (bezw m n) * int n = int (gcd m n)"
- apply (simp add: bezw_non_0 nat_gcd_non_0)
+ apply (simp add: bezw_non_0 gcd_non_0_nat)
apply (erule subst)
apply (simp add: ring_simps)
apply (subst mod_div_equality [of m n, symmetric])
@@ -1075,7 +1075,7 @@
done
qed
-lemma int_bezout:
+lemma bezout_int:
fixes x y
shows "EX u v. u * (x::int) + v * y = gcd x y"
proof -
@@ -1098,7 +1098,7 @@
apply auto
apply (rule_tac x = u in exI)
apply (rule_tac x = "-v" in exI)
- apply (subst int_gcd_neg2 [symmetric])
+ apply (subst gcd_neg2_int [symmetric])
apply auto
done
moreover have "x <= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> ?thesis"
@@ -1106,7 +1106,7 @@
apply auto
apply (rule_tac x = "-u" in exI)
apply (rule_tac x = v in exI)
- apply (subst int_gcd_neg1 [symmetric])
+ apply (subst gcd_neg1_int [symmetric])
apply auto
done
moreover have "x <= 0 \<Longrightarrow> y <= 0 \<Longrightarrow> ?thesis"
@@ -1114,8 +1114,8 @@
apply auto
apply (rule_tac x = "-u" in exI)
apply (rule_tac x = "-v" in exI)
- apply (subst int_gcd_neg1 [symmetric])
- apply (subst int_gcd_neg2 [symmetric])
+ apply (subst gcd_neg1_int [symmetric])
+ apply (subst gcd_neg2_int [symmetric])
apply auto
done
ultimately show ?thesis by blast
@@ -1160,7 +1160,7 @@
ultimately show "P a b" by blast
qed
-lemma nat_bezout_lemma:
+lemma bezout_lemma_nat:
assumes ex: "\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and>
(a * x = b * y + d \<or> b * x = a * y + d)"
shows "\<exists>d x y. d dvd a \<and> d dvd a + b \<and>
@@ -1177,7 +1177,7 @@
apply algebra
done
-lemma nat_bezout_add: "\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and>
+lemma bezout_add_nat: "\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and>
(a * x = b * y + d \<or> b * x = a * y + d)"
apply(induct a b rule: ind_euclid)
apply blast
@@ -1194,9 +1194,9 @@
apply algebra
done
-lemma nat_bezout1: "\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and>
+lemma bezout1_nat: "\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and>
(a * x - b * y = d \<or> b * x - a * y = d)"
- using nat_bezout_add[of a b]
+ using bezout_add_nat[of a b]
apply clarsimp
apply (rule_tac x="d" in exI, simp)
apply (rule_tac x="x" in exI)
@@ -1204,11 +1204,11 @@
apply auto
done
-lemma nat_bezout_add_strong: assumes nz: "a \<noteq> (0::nat)"
+lemma bezout_add_strong_nat: assumes nz: "a \<noteq> (0::nat)"
shows "\<exists>d x y. d dvd a \<and> d dvd b \<and> a * x = b * y + d"
proof-
from nz have ap: "a > 0" by simp
- from nat_bezout_add[of a b]
+ from bezout_add_nat[of a b]
have "(\<exists>d x y. d dvd a \<and> d dvd b \<and> a * x = b * y + d) \<or>
(\<exists>d x y. d dvd a \<and> d dvd b \<and> b * x = a * y + d)" by blast
moreover
@@ -1258,11 +1258,11 @@
ultimately show ?thesis by blast
qed
-lemma nat_bezout: assumes a: "(a::nat) \<noteq> 0"
+lemma bezout_nat: assumes a: "(a::nat) \<noteq> 0"
shows "\<exists>x y. a * x = b * y + gcd a b"
proof-
let ?g = "gcd a b"
- from nat_bezout_add_strong[OF a, of b]
+ from bezout_add_strong_nat[OF a, of b]
obtain d x y where d: "d dvd a" "d dvd b" "a * x = b * y + d" by blast
from d(1,2) have "d dvd ?g" by simp
then obtain k where k: "?g = d*k" unfolding dvd_def by blast
@@ -1274,99 +1274,99 @@
subsection {* LCM *}
-lemma int_lcm_altdef: "lcm (a::int) b = (abs a) * (abs b) div gcd a b"
+lemma lcm_altdef_int: "lcm (a::int) b = (abs a) * (abs b) div gcd a b"
by (simp add: lcm_int_def lcm_nat_def zdiv_int
zmult_int [symmetric] gcd_int_def)
-lemma nat_prod_gcd_lcm: "(m::nat) * n = gcd m n * lcm m n"
+lemma prod_gcd_lcm_nat: "(m::nat) * n = gcd m n * lcm m n"
unfolding lcm_nat_def
- by (simp add: dvd_mult_div_cancel [OF nat_gcd_dvd_prod])
+ by (simp add: dvd_mult_div_cancel [OF gcd_dvd_prod_nat])
-lemma int_prod_gcd_lcm: "abs(m::int) * abs n = gcd m n * lcm m n"
+lemma prod_gcd_lcm_int: "abs(m::int) * abs n = gcd m n * lcm m n"
unfolding lcm_int_def gcd_int_def
apply (subst int_mult [symmetric])
- apply (subst nat_prod_gcd_lcm [symmetric])
+ apply (subst prod_gcd_lcm_nat [symmetric])
apply (subst nat_abs_mult_distrib [symmetric])
apply (simp, simp add: abs_mult)
done
-lemma nat_lcm_0 [simp]: "lcm (m::nat) 0 = 0"
+lemma lcm_0_nat [simp]: "lcm (m::nat) 0 = 0"
unfolding lcm_nat_def by simp
-lemma int_lcm_0 [simp]: "lcm (m::int) 0 = 0"
+lemma lcm_0_int [simp]: "lcm (m::int) 0 = 0"
unfolding lcm_int_def by simp
-lemma nat_lcm_0_left [simp]: "lcm (0::nat) n = 0"
+lemma lcm_0_left_nat [simp]: "lcm (0::nat) n = 0"
unfolding lcm_nat_def by simp
-lemma int_lcm_0_left [simp]: "lcm (0::int) n = 0"
+lemma lcm_0_left_int [simp]: "lcm (0::int) n = 0"
unfolding lcm_int_def by simp
-lemma nat_lcm_commute: "lcm (m::nat) n = lcm n m"
- unfolding lcm_nat_def by (simp add: nat_gcd_commute ring_simps)
+lemma lcm_commute_nat: "lcm (m::nat) n = lcm n m"
+ unfolding lcm_nat_def by (simp add: gcd_commute_nat ring_simps)
-lemma int_lcm_commute: "lcm (m::int) n = lcm n m"
- unfolding lcm_int_def by (subst nat_lcm_commute, rule refl)
+lemma lcm_commute_int: "lcm (m::int) n = lcm n m"
+ unfolding lcm_int_def by (subst lcm_commute_nat, rule refl)
-lemma nat_lcm_pos:
+lemma lcm_pos_nat:
"(m::nat) > 0 \<Longrightarrow> n>0 \<Longrightarrow> lcm m n > 0"
-by (metis gr0I mult_is_0 nat_prod_gcd_lcm)
+by (metis gr0I mult_is_0 prod_gcd_lcm_nat)
-lemma int_lcm_pos:
+lemma lcm_pos_int:
"(m::int) ~= 0 \<Longrightarrow> n ~= 0 \<Longrightarrow> lcm m n > 0"
- apply (subst int_lcm_abs)
- apply (rule nat_lcm_pos [transferred])
+ apply (subst lcm_abs_int)
+ apply (rule lcm_pos_nat [transferred])
apply auto
done
-lemma nat_dvd_pos:
+lemma dvd_pos_nat:
fixes n m :: nat
assumes "n > 0" and "m dvd n"
shows "m > 0"
using assms by (cases m) auto
-lemma nat_lcm_least:
+lemma lcm_least_nat:
assumes "(m::nat) dvd k" and "n dvd k"
shows "lcm m n dvd k"
proof (cases k)
case 0 then show ?thesis by auto
next
case (Suc _) then have pos_k: "k > 0" by auto
- from assms nat_dvd_pos [OF this] have pos_mn: "m > 0" "n > 0" by auto
- with nat_gcd_zero [of m n] have pos_gcd: "gcd m n > 0" by simp
+ from assms dvd_pos_nat [OF this] have pos_mn: "m > 0" "n > 0" by auto
+ with gcd_zero_nat [of m n] have pos_gcd: "gcd m n > 0" by simp
from assms obtain p where k_m: "k = m * p" using dvd_def by blast
from assms obtain q where k_n: "k = n * q" using dvd_def by blast
from pos_k k_m have pos_p: "p > 0" by auto
from pos_k k_n have pos_q: "q > 0" by auto
have "k * k * gcd q p = k * gcd (k * q) (k * p)"
- by (simp add: mult_ac nat_gcd_mult_distrib)
+ by (simp add: mult_ac gcd_mult_distrib_nat)
also have "\<dots> = k * gcd (m * p * q) (n * q * p)"
by (simp add: k_m [symmetric] k_n [symmetric])
also have "\<dots> = k * p * q * gcd m n"
- by (simp add: mult_ac nat_gcd_mult_distrib)
+ by (simp add: mult_ac gcd_mult_distrib_nat)
finally have "(m * p) * (n * q) * gcd q p = k * p * q * gcd m n"
by (simp only: k_m [symmetric] k_n [symmetric])
then have "p * q * m * n * gcd q p = p * q * k * gcd m n"
by (simp add: mult_ac)
with pos_p pos_q have "m * n * gcd q p = k * gcd m n"
by simp
- with nat_prod_gcd_lcm [of m n]
+ with prod_gcd_lcm_nat [of m n]
have "lcm m n * gcd q p * gcd m n = k * gcd m n"
by (simp add: mult_ac)
with pos_gcd have "lcm m n * gcd q p = k" by auto
then show ?thesis using dvd_def by auto
qed
-lemma int_lcm_least:
+lemma lcm_least_int:
"(m::int) dvd k \<Longrightarrow> n dvd k \<Longrightarrow> lcm m n dvd k"
-apply (subst int_lcm_abs)
+apply (subst lcm_abs_int)
apply (rule dvd_trans)
-apply (rule nat_lcm_least [transferred, of _ "abs k" _])
+apply (rule lcm_least_nat [transferred, of _ "abs k" _])
apply auto
done
-lemma nat_lcm_dvd1: "(m::nat) dvd lcm m n"
+lemma lcm_dvd1_nat: "(m::nat) dvd lcm m n"
proof (cases m)
case 0 then show ?thesis by simp
next
@@ -1382,82 +1382,82 @@
then obtain k where "n = gcd m n * k" using dvd_def by auto
then have "m * n div gcd m n = m * (gcd m n * k) div gcd m n"
by (simp add: mult_ac)
- also have "\<dots> = m * k" using mpos npos nat_gcd_zero by simp
+ also have "\<dots> = m * k" using mpos npos gcd_zero_nat by simp
finally show ?thesis by (simp add: lcm_nat_def)
qed
qed
-lemma int_lcm_dvd1: "(m::int) dvd lcm m n"
- apply (subst int_lcm_abs)
+lemma lcm_dvd1_int: "(m::int) dvd lcm m n"
+ apply (subst lcm_abs_int)
apply (rule dvd_trans)
prefer 2
- apply (rule nat_lcm_dvd1 [transferred])
+ apply (rule lcm_dvd1_nat [transferred])
apply auto
done
-lemma nat_lcm_dvd2: "(n::nat) dvd lcm m n"
- by (subst nat_lcm_commute, rule nat_lcm_dvd1)
+lemma lcm_dvd2_nat: "(n::nat) dvd lcm m n"
+ by (subst lcm_commute_nat, rule lcm_dvd1_nat)
-lemma int_lcm_dvd2: "(n::int) dvd lcm m n"
- by (subst int_lcm_commute, rule int_lcm_dvd1)
+lemma lcm_dvd2_int: "(n::int) dvd lcm m n"
+ by (subst lcm_commute_int, rule lcm_dvd1_int)
lemma dvd_lcm_I1_nat[simp]: "(k::nat) dvd m \<Longrightarrow> k dvd lcm m n"
-by(metis nat_lcm_dvd1 dvd_trans)
+by(metis lcm_dvd1_nat dvd_trans)
lemma dvd_lcm_I2_nat[simp]: "(k::nat) dvd n \<Longrightarrow> k dvd lcm m n"
-by(metis nat_lcm_dvd2 dvd_trans)
+by(metis lcm_dvd2_nat dvd_trans)
lemma dvd_lcm_I1_int[simp]: "(i::int) dvd m \<Longrightarrow> i dvd lcm m n"
-by(metis int_lcm_dvd1 dvd_trans)
+by(metis lcm_dvd1_int dvd_trans)
lemma dvd_lcm_I2_int[simp]: "(i::int) dvd n \<Longrightarrow> i dvd lcm m n"
-by(metis int_lcm_dvd2 dvd_trans)
+by(metis lcm_dvd2_int dvd_trans)
-lemma nat_lcm_unique: "(a::nat) dvd d \<and> b dvd d \<and>
+lemma lcm_unique_nat: "(a::nat) dvd d \<and> b dvd d \<and>
(\<forall>e. a dvd e \<and> b dvd e \<longrightarrow> d dvd e) \<longleftrightarrow> d = lcm a b"
- by (auto intro: dvd_anti_sym nat_lcm_least nat_lcm_dvd1 nat_lcm_dvd2)
+ by (auto intro: dvd_anti_sym lcm_least_nat lcm_dvd1_nat lcm_dvd2_nat)
-lemma int_lcm_unique: "d >= 0 \<and> (a::int) dvd d \<and> b dvd d \<and>
+lemma lcm_unique_int: "d >= 0 \<and> (a::int) dvd d \<and> b dvd d \<and>
(\<forall>e. a dvd e \<and> b dvd e \<longrightarrow> d dvd e) \<longleftrightarrow> d = lcm a b"
- by (auto intro: dvd_anti_sym [transferred] int_lcm_least)
+ by (auto intro: dvd_anti_sym [transferred] lcm_least_int)
lemma lcm_proj2_if_dvd_nat [simp]: "(x::nat) dvd y \<Longrightarrow> lcm x y = y"
apply (rule sym)
- apply (subst nat_lcm_unique [symmetric])
+ apply (subst lcm_unique_nat [symmetric])
apply auto
done
lemma lcm_proj2_if_dvd_int [simp]: "(x::int) dvd y \<Longrightarrow> lcm x y = abs y"
apply (rule sym)
- apply (subst int_lcm_unique [symmetric])
+ apply (subst lcm_unique_int [symmetric])
apply auto
done
lemma lcm_proj1_if_dvd_nat [simp]: "(x::nat) dvd y \<Longrightarrow> lcm y x = y"
-by (subst nat_lcm_commute, erule lcm_proj2_if_dvd_nat)
+by (subst lcm_commute_nat, erule lcm_proj2_if_dvd_nat)
lemma lcm_proj1_if_dvd_int [simp]: "(x::int) dvd y \<Longrightarrow> lcm y x = abs y"
-by (subst int_lcm_commute, erule lcm_proj2_if_dvd_int)
+by (subst lcm_commute_int, erule lcm_proj2_if_dvd_int)
lemma lcm_assoc_nat: "lcm (lcm n m) (p::nat) = lcm n (lcm m p)"
-apply(rule nat_lcm_unique[THEN iffD1])
-apply (metis dvd.order_trans nat_lcm_unique)
+apply(rule lcm_unique_nat[THEN iffD1])
+apply (metis dvd.order_trans lcm_unique_nat)
done
lemma lcm_assoc_int: "lcm (lcm n m) (p::int) = lcm n (lcm m p)"
-apply(rule int_lcm_unique[THEN iffD1])
-apply (metis dvd_trans int_lcm_unique)
+apply(rule lcm_unique_int[THEN iffD1])
+apply (metis dvd_trans lcm_unique_int)
done
lemmas lcm_left_commute_nat =
- mk_left_commute[of lcm, OF lcm_assoc_nat nat_lcm_commute]
+ mk_left_commute[of lcm, OF lcm_assoc_nat lcm_commute_nat]
lemmas lcm_left_commute_int =
- mk_left_commute[of lcm, OF lcm_assoc_int int_lcm_commute]
+ mk_left_commute[of lcm, OF lcm_assoc_int lcm_commute_int]
-lemmas lcm_ac_nat = lcm_assoc_nat nat_lcm_commute lcm_left_commute_nat
-lemmas lcm_ac_int = lcm_assoc_int int_lcm_commute lcm_left_commute_int
+lemmas lcm_ac_nat = lcm_assoc_nat lcm_commute_nat lcm_left_commute_nat
+lemmas lcm_ac_int = lcm_assoc_int lcm_commute_int lcm_left_commute_int
subsection {* Primes *}
@@ -1465,40 +1465,40 @@
(* Is there a better way to handle these, rather than making them
elim rules? *)
-lemma nat_prime_ge_0 [elim]: "prime (p::nat) \<Longrightarrow> p >= 0"
+lemma prime_ge_0_nat [elim]: "prime (p::nat) \<Longrightarrow> p >= 0"
by (unfold prime_nat_def, auto)
-lemma nat_prime_gt_0 [elim]: "prime (p::nat) \<Longrightarrow> p > 0"
+lemma prime_gt_0_nat [elim]: "prime (p::nat) \<Longrightarrow> p > 0"
by (unfold prime_nat_def, auto)
-lemma nat_prime_ge_1 [elim]: "prime (p::nat) \<Longrightarrow> p >= 1"
+lemma prime_ge_1_nat [elim]: "prime (p::nat) \<Longrightarrow> p >= 1"
by (unfold prime_nat_def, auto)
-lemma nat_prime_gt_1 [elim]: "prime (p::nat) \<Longrightarrow> p > 1"
+lemma prime_gt_1_nat [elim]: "prime (p::nat) \<Longrightarrow> p > 1"
by (unfold prime_nat_def, auto)
-lemma nat_prime_ge_Suc_0 [elim]: "prime (p::nat) \<Longrightarrow> p >= Suc 0"
+lemma prime_ge_Suc_0_nat [elim]: "prime (p::nat) \<Longrightarrow> p >= Suc 0"
by (unfold prime_nat_def, auto)
-lemma nat_prime_gt_Suc_0 [elim]: "prime (p::nat) \<Longrightarrow> p > Suc 0"
+lemma prime_gt_Suc_0_nat [elim]: "prime (p::nat) \<Longrightarrow> p > Suc 0"
by (unfold prime_nat_def, auto)
-lemma nat_prime_ge_2 [elim]: "prime (p::nat) \<Longrightarrow> p >= 2"
+lemma prime_ge_2_nat [elim]: "prime (p::nat) \<Longrightarrow> p >= 2"
by (unfold prime_nat_def, auto)
-lemma int_prime_ge_0 [elim]: "prime (p::int) \<Longrightarrow> p >= 0"
+lemma prime_ge_0_int [elim]: "prime (p::int) \<Longrightarrow> p >= 0"
by (unfold prime_int_def prime_nat_def, auto)
-lemma int_prime_gt_0 [elim]: "prime (p::int) \<Longrightarrow> p > 0"
+lemma prime_gt_0_int [elim]: "prime (p::int) \<Longrightarrow> p > 0"
by (unfold prime_int_def prime_nat_def, auto)
-lemma int_prime_ge_1 [elim]: "prime (p::int) \<Longrightarrow> p >= 1"
+lemma prime_ge_1_int [elim]: "prime (p::int) \<Longrightarrow> p >= 1"
by (unfold prime_int_def prime_nat_def, auto)
-lemma int_prime_gt_1 [elim]: "prime (p::int) \<Longrightarrow> p > 1"
+lemma prime_gt_1_int [elim]: "prime (p::int) \<Longrightarrow> p > 1"
by (unfold prime_int_def prime_nat_def, auto)
-lemma int_prime_ge_2 [elim]: "prime (p::int) \<Longrightarrow> p >= 2"
+lemma prime_ge_2_int [elim]: "prime (p::int) \<Longrightarrow> p >= 2"
by (unfold prime_int_def prime_nat_def, auto)
thm prime_nat_def;
@@ -1508,59 +1508,59 @@
m = 1 \<or> m = p))"
using prime_nat_def [transferred]
apply (case_tac "p >= 0")
- by (blast, auto simp add: int_prime_ge_0)
+ by (blast, auto simp add: prime_ge_0_int)
(* To do: determine primality of any numeral *)
-lemma nat_zero_not_prime [simp]: "~prime (0::nat)"
+lemma zero_not_prime_nat [simp]: "~prime (0::nat)"
by (simp add: prime_nat_def)
-lemma int_zero_not_prime [simp]: "~prime (0::int)"
+lemma zero_not_prime_int [simp]: "~prime (0::int)"
by (simp add: prime_int_def)
-lemma nat_one_not_prime [simp]: "~prime (1::nat)"
+lemma one_not_prime_nat [simp]: "~prime (1::nat)"
by (simp add: prime_nat_def)
-lemma nat_Suc_0_not_prime [simp]: "~prime (Suc 0)"
+lemma Suc_0_not_prime_nat [simp]: "~prime (Suc 0)"
by (simp add: prime_nat_def One_nat_def)
-lemma int_one_not_prime [simp]: "~prime (1::int)"
+lemma one_not_prime_int [simp]: "~prime (1::int)"
by (simp add: prime_int_def)
-lemma nat_two_is_prime [simp]: "prime (2::nat)"
+lemma two_is_prime_nat [simp]: "prime (2::nat)"
apply (auto simp add: prime_nat_def)
apply (case_tac m)
apply (auto dest!: dvd_imp_le)
done
-lemma int_two_is_prime [simp]: "prime (2::int)"
- by (rule nat_two_is_prime [transferred direction: nat "op <= (0::int)"])
+lemma two_is_prime_int [simp]: "prime (2::int)"
+ by (rule two_is_prime_nat [transferred direction: nat "op <= (0::int)"])
-lemma nat_prime_imp_coprime: "prime (p::nat) \<Longrightarrow> \<not> p dvd n \<Longrightarrow> coprime p n"
+lemma prime_imp_coprime_nat: "prime (p::nat) \<Longrightarrow> \<not> p dvd n \<Longrightarrow> coprime p n"
apply (unfold prime_nat_def)
- apply (metis nat_gcd_dvd1 nat_gcd_dvd2)
+ apply (metis gcd_dvd1_nat gcd_dvd2_nat)
done
-lemma int_prime_imp_coprime: "prime (p::int) \<Longrightarrow> \<not> p dvd n \<Longrightarrow> coprime p n"
+lemma prime_imp_coprime_int: "prime (p::int) \<Longrightarrow> \<not> p dvd n \<Longrightarrow> coprime p n"
apply (unfold prime_int_altdef)
- apply (metis int_gcd_dvd1 int_gcd_dvd2 int_gcd_ge_0)
+ apply (metis gcd_dvd1_int gcd_dvd2_int gcd_ge_0_int)
done
-lemma nat_prime_dvd_mult: "prime (p::nat) \<Longrightarrow> p dvd m * n \<Longrightarrow> p dvd m \<or> p dvd n"
- by (blast intro: nat_coprime_dvd_mult nat_prime_imp_coprime)
+lemma prime_dvd_mult_nat: "prime (p::nat) \<Longrightarrow> p dvd m * n \<Longrightarrow> p dvd m \<or> p dvd n"
+ by (blast intro: coprime_dvd_mult_nat prime_imp_coprime_nat)
-lemma int_prime_dvd_mult: "prime (p::int) \<Longrightarrow> p dvd m * n \<Longrightarrow> p dvd m \<or> p dvd n"
- by (blast intro: int_coprime_dvd_mult int_prime_imp_coprime)
+lemma prime_dvd_mult_int: "prime (p::int) \<Longrightarrow> p dvd m * n \<Longrightarrow> p dvd m \<or> p dvd n"
+ by (blast intro: coprime_dvd_mult_int prime_imp_coprime_int)
-lemma nat_prime_dvd_mult_eq [simp]: "prime (p::nat) \<Longrightarrow>
+lemma prime_dvd_mult_eq_nat [simp]: "prime (p::nat) \<Longrightarrow>
p dvd m * n = (p dvd m \<or> p dvd n)"
- by (rule iffI, rule nat_prime_dvd_mult, auto)
+ by (rule iffI, rule prime_dvd_mult_nat, auto)
-lemma int_prime_dvd_mult_eq [simp]: "prime (p::int) \<Longrightarrow>
+lemma prime_dvd_mult_eq_int [simp]: "prime (p::int) \<Longrightarrow>
p dvd m * n = (p dvd m \<or> p dvd n)"
- by (rule iffI, rule int_prime_dvd_mult, auto)
+ by (rule iffI, rule prime_dvd_mult_int, auto)
-lemma nat_not_prime_eq_prod: "(n::nat) > 1 \<Longrightarrow> ~ prime n \<Longrightarrow>
+lemma not_prime_eq_prod_nat: "(n::nat) > 1 \<Longrightarrow> ~ prime n \<Longrightarrow>
EX m k. n = m * k & 1 < m & m < n & 1 < k & k < n"
unfolding prime_nat_def dvd_def apply auto
apply (subgoal_tac "k > 1")
@@ -1573,7 +1573,7 @@
(* there's a lot of messing around with signs of products here --
could this be made more automatic? *)
-lemma int_not_prime_eq_prod: "(n::int) > 1 \<Longrightarrow> ~ prime n \<Longrightarrow>
+lemma not_prime_eq_prod_int: "(n::int) > 1 \<Longrightarrow> ~ prime n \<Longrightarrow>
EX m k. n = m * k & 1 < m & m < n & 1 < k & k < n"
unfolding prime_int_altdef dvd_def
apply auto
@@ -1593,72 +1593,72 @@
apply auto
done
-lemma nat_prime_dvd_power [rule_format]: "prime (p::nat) -->
+lemma prime_dvd_power_nat [rule_format]: "prime (p::nat) -->
n > 0 --> (p dvd x^n --> p dvd x)"
by (induct n rule: nat_induct, auto)
-lemma int_prime_dvd_power [rule_format]: "prime (p::int) -->
+lemma prime_dvd_power_int [rule_format]: "prime (p::int) -->
n > 0 --> (p dvd x^n --> p dvd x)"
apply (induct n rule: nat_induct, auto)
- apply (frule int_prime_ge_0)
+ apply (frule prime_ge_0_int)
apply auto
done
-lemma nat_prime_imp_power_coprime: "prime (p::nat) \<Longrightarrow> ~ p dvd a \<Longrightarrow>
+lemma prime_imp_power_coprime_nat: "prime (p::nat) \<Longrightarrow> ~ p dvd a \<Longrightarrow>
coprime a (p^m)"
- apply (rule nat_coprime_exp)
- apply (subst nat_gcd_commute)
- apply (erule (1) nat_prime_imp_coprime)
+ apply (rule coprime_exp_nat)
+ apply (subst gcd_commute_nat)
+ apply (erule (1) prime_imp_coprime_nat)
done
-lemma int_prime_imp_power_coprime: "prime (p::int) \<Longrightarrow> ~ p dvd a \<Longrightarrow>
+lemma prime_imp_power_coprime_int: "prime (p::int) \<Longrightarrow> ~ p dvd a \<Longrightarrow>
coprime a (p^m)"
- apply (rule int_coprime_exp)
- apply (subst int_gcd_commute)
- apply (erule (1) int_prime_imp_coprime)
+ apply (rule coprime_exp_int)
+ apply (subst gcd_commute_int)
+ apply (erule (1) prime_imp_coprime_int)
done
-lemma nat_primes_coprime: "prime (p::nat) \<Longrightarrow> prime q \<Longrightarrow> p \<noteq> q \<Longrightarrow> coprime p q"
- apply (rule nat_prime_imp_coprime, assumption)
+lemma primes_coprime_nat: "prime (p::nat) \<Longrightarrow> prime q \<Longrightarrow> p \<noteq> q \<Longrightarrow> coprime p q"
+ apply (rule prime_imp_coprime_nat, assumption)
apply (unfold prime_nat_def, auto)
done
-lemma int_primes_coprime: "prime (p::int) \<Longrightarrow> prime q \<Longrightarrow> p \<noteq> q \<Longrightarrow> coprime p q"
- apply (rule int_prime_imp_coprime, assumption)
+lemma primes_coprime_int: "prime (p::int) \<Longrightarrow> prime q \<Longrightarrow> p \<noteq> q \<Longrightarrow> coprime p q"
+ apply (rule prime_imp_coprime_int, assumption)
apply (unfold prime_int_altdef, clarify)
apply (drule_tac x = q in spec)
apply (drule_tac x = p in spec)
apply auto
done
-lemma nat_primes_imp_powers_coprime: "prime (p::nat) \<Longrightarrow> prime q \<Longrightarrow> p ~= q \<Longrightarrow>
+lemma primes_imp_powers_coprime_nat: "prime (p::nat) \<Longrightarrow> prime q \<Longrightarrow> p ~= q \<Longrightarrow>
coprime (p^m) (q^n)"
- by (rule nat_coprime_exp2, rule nat_primes_coprime)
+ by (rule coprime_exp2_nat, rule primes_coprime_nat)
-lemma int_primes_imp_powers_coprime: "prime (p::int) \<Longrightarrow> prime q \<Longrightarrow> p ~= q \<Longrightarrow>
+lemma primes_imp_powers_coprime_int: "prime (p::int) \<Longrightarrow> prime q \<Longrightarrow> p ~= q \<Longrightarrow>
coprime (p^m) (q^n)"
- by (rule int_coprime_exp2, rule int_primes_coprime)
+ by (rule coprime_exp2_int, rule primes_coprime_int)
-lemma nat_prime_factor: "n \<noteq> (1::nat) \<Longrightarrow> \<exists> p. prime p \<and> p dvd n"
+lemma prime_factor_nat: "n \<noteq> (1::nat) \<Longrightarrow> \<exists> p. prime p \<and> p dvd n"
apply (induct n rule: nat_less_induct)
apply (case_tac "n = 0")
- using nat_two_is_prime apply blast
+ using two_is_prime_nat apply blast
apply (case_tac "prime n")
apply blast
apply (subgoal_tac "n > 1")
- apply (frule (1) nat_not_prime_eq_prod)
+ apply (frule (1) not_prime_eq_prod_nat)
apply (auto intro: dvd_mult dvd_mult2)
done
(* An Isar version:
-lemma nat_prime_factor_b:
+lemma prime_factor_b_nat:
fixes n :: nat
assumes "n \<noteq> 1"
shows "\<exists>p. prime p \<and> p dvd n"
using `n ~= 1`
-proof (induct n rule: nat_less_induct)
+proof (induct n rule: less_induct_nat)
fix n :: nat
assume "n ~= 1" and
ih: "\<forall>m<n. m \<noteq> 1 \<longrightarrow> (\<exists>p. prime p \<and> p dvd m)"
@@ -1666,9 +1666,9 @@
proof -
{
assume "n = 0"
- moreover note nat_two_is_prime
+ moreover note two_is_prime_nat
ultimately have ?thesis
- by (auto simp del: nat_two_is_prime)
+ by (auto simp del: two_is_prime_nat)
}
moreover
{
@@ -1679,7 +1679,7 @@
{
assume "n ~= 0" and "~ prime n"
with `n ~= 1` have "n > 1" by auto
- with `~ prime n` and nat_not_prime_eq_prod obtain m k where
+ with `~ prime n` and not_prime_eq_prod_nat obtain m k where
"n = m * k" and "1 < m" and "m < n" by blast
with ih obtain p where "prime p" and "p dvd m" by blast
with `n = m * k` have ?thesis by auto
@@ -1692,7 +1692,7 @@
text {* One property of coprimality is easier to prove via prime factors. *}
-lemma nat_prime_divprod_pow:
+lemma prime_divprod_pow_nat:
assumes p: "prime (p::nat)" and ab: "coprime a b" and pab: "p^n dvd a * b"
shows "p^n dvd a \<or> p^n dvd b"
proof-
@@ -1705,27 +1705,27 @@
from n have "p dvd p^n" by (intro dvd_power, auto)
also note pab
finally have pab': "p dvd a * b".
- from nat_prime_dvd_mult[OF p pab']
+ from prime_dvd_mult_nat[OF p pab']
have "p dvd a \<or> p dvd b" .
moreover
{assume pa: "p dvd a"
have pnba: "p^n dvd b*a" using pab by (simp add: mult_commute)
- from nat_coprime_common_divisor [OF ab, OF pa] p have "\<not> p dvd b" by auto
+ from coprime_common_divisor_nat [OF ab, OF pa] p have "\<not> p dvd b" by auto
with p have "coprime b p"
- by (subst nat_gcd_commute, intro nat_prime_imp_coprime)
+ by (subst gcd_commute_nat, intro prime_imp_coprime_nat)
hence pnb: "coprime (p^n) b"
- by (subst nat_gcd_commute, rule nat_coprime_exp)
- from nat_coprime_divprod[OF pnba pnb] have ?thesis by blast }
+ by (subst gcd_commute_nat, rule coprime_exp_nat)
+ from coprime_divprod_nat[OF pnba pnb] have ?thesis by blast }
moreover
{assume pb: "p dvd b"
have pnba: "p^n dvd b*a" using pab by (simp add: mult_commute)
- from nat_coprime_common_divisor [OF ab, of p] pb p have "\<not> p dvd a"
+ from coprime_common_divisor_nat [OF ab, of p] pb p have "\<not> p dvd a"
by auto
with p have "coprime a p"
- by (subst nat_gcd_commute, intro nat_prime_imp_coprime)
+ by (subst gcd_commute_nat, intro prime_imp_coprime_nat)
hence pna: "coprime (p^n) a"
- by (subst nat_gcd_commute, rule nat_coprime_exp)
- from nat_coprime_divprod[OF pab pna] have ?thesis by blast }
+ by (subst gcd_commute_nat, rule coprime_exp_nat)
+ from coprime_divprod_nat[OF pab pna] have ?thesis by blast }
ultimately have ?thesis by blast}
ultimately show ?thesis by blast
qed
--- a/src/HOL/HOL.thy Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/HOL.thy Thu Jul 09 08:55:42 2009 +0200
@@ -1869,7 +1869,27 @@
declare simp_thms(6) [code nbe]
setup {*
- Code.add_const_alias @{thm equals_eq}
+ Sign.add_const_constraint (@{const_name eq}, SOME @{typ "'a\<Colon>type \<Rightarrow> 'a \<Rightarrow> bool"})
+*}
+
+lemma equals_alias_cert: "OFCLASS('a, eq_class) \<equiv> ((op = :: 'a \<Rightarrow> 'a \<Rightarrow> bool) \<equiv> eq)" (is "?ofclass \<equiv> ?eq")
+proof
+ assume "PROP ?ofclass"
+ show "PROP ?eq"
+ by (tactic {* ALLGOALS (rtac (Drule.unconstrainTs @{thm equals_eq})) *})
+ (fact `PROP ?ofclass`)
+next
+ assume "PROP ?eq"
+ show "PROP ?ofclass" proof
+ qed (simp add: `PROP ?eq`)
+qed
+
+setup {*
+ Sign.add_const_constraint (@{const_name eq}, SOME @{typ "'a\<Colon>eq \<Rightarrow> 'a \<Rightarrow> bool"})
+*}
+
+setup {*
+ Code.add_const_alias @{thm equals_alias_cert}
*}
hide (open) const eq
--- a/src/HOL/Import/import.ML Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/Import/import.ML Thu Jul 09 08:55:42 2009 +0200
@@ -55,7 +55,7 @@
then message "import: Terms match up"
else message "import: Terms DO NOT match up"
val thm' = equal_elim (symmetric prew) thm
- val res = bicompose true (false,thm',0) 1 th
+ val res = Thm.bicompose true (false,thm',0) 1 th
in
res
end
--- a/src/HOL/Import/proof_kernel.ML Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/Import/proof_kernel.ML Thu Jul 09 08:55:42 2009 +0200
@@ -977,7 +977,7 @@
fun uniq_compose m th i st =
let
- val res = bicompose false (false,th,m) i st
+ val res = Thm.bicompose false (false,th,m) i st
in
case Seq.pull res of
SOME (th,rest) => (case Seq.pull rest of
@@ -1065,14 +1065,12 @@
res
end
-val permute_prems = Thm.permute_prems
-
fun rearrange sg tm th =
let
val tm' = Envir.beta_eta_contract tm
- fun find [] n = permute_prems 0 1 (implies_intr (Thm.cterm_of sg tm) th)
+ fun find [] n = Thm.permute_prems 0 1 (implies_intr (Thm.cterm_of sg tm) th)
| find (p::ps) n = if tm' aconv (Envir.beta_eta_contract p)
- then permute_prems n 1 th
+ then Thm.permute_prems n 1 th
else find ps (n+1)
in
find (prems_of th) 0
@@ -1193,7 +1191,7 @@
fun if_debug f x = if !debug then f x else ()
val message = if_debug writeln
-val conjE_helper = permute_prems 0 1 conjE
+val conjE_helper = Thm.permute_prems 0 1 conjE
fun get_hol4_thm thyname thmname thy =
case get_hol4_theorem thyname thmname thy of
--- a/src/HOL/Import/shuffler.ML Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/Import/shuffler.ML Thu Jul 09 08:55:42 2009 +0200
@@ -595,7 +595,7 @@
val norm_th = Thm.varifyT (norm_thm thy (close_thm (Thm.transfer thy th)))
val triv_th = trivial rhs
val _ = message ("Shuffler.set_prop: Gluing together " ^ (string_of_thm norm_th) ^ " and " ^ (string_of_thm triv_th))
- val mod_th = case Seq.pull (bicompose false (*true*) (false,norm_th,0) 1 triv_th) of
+ val mod_th = case Seq.pull (Thm.bicompose false (*true*) (false,norm_th,0) 1 triv_th) of
SOME(th,_) => SOME th
| NONE => NONE
in
--- a/src/HOL/Inductive.thy Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/Inductive.thy Thu Jul 09 08:55:42 2009 +0200
@@ -258,38 +258,6 @@
subsection {* Inductive predicates and sets *}
-text {* Inversion of injective functions. *}
-
-constdefs
- myinv :: "('a => 'b) => ('b => 'a)"
- "myinv (f :: 'a => 'b) == \<lambda>y. THE x. f x = y"
-
-lemma myinv_f_f: "inj f ==> myinv f (f x) = x"
-proof -
- assume "inj f"
- hence "(THE x'. f x' = f x) = (THE x'. x' = x)"
- by (simp only: inj_eq)
- also have "... = x" by (rule the_eq_trivial)
- finally show ?thesis by (unfold myinv_def)
-qed
-
-lemma f_myinv_f: "inj f ==> y \<in> range f ==> f (myinv f y) = y"
-proof (unfold myinv_def)
- assume inj: "inj f"
- assume "y \<in> range f"
- then obtain x where "y = f x" ..
- hence x: "f x = y" ..
- thus "f (THE x. f x = y) = y"
- proof (rule theI)
- fix x' assume "f x' = y"
- with x have "f x' = f x" by simp
- with inj show "x' = x" by (rule injD)
- qed
-qed
-
-hide const myinv
-
-
text {* Package setup. *}
theorems basic_monos =
--- a/src/HOL/Library/Abstract_Rat.thy Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/Library/Abstract_Rat.thy Thu Jul 09 08:55:42 2009 +0200
@@ -30,8 +30,8 @@
(let g = gcd a b
in if b > 0 then (a div g, b div g) else (- (a div g), - (b div g)))))"
-declare int_gcd_dvd1[presburger]
-declare int_gcd_dvd2[presburger]
+declare gcd_dvd1_int[presburger]
+declare gcd_dvd2_int[presburger]
lemma normNum_isnormNum [simp]: "isnormNum (normNum x)"
proof -
have " \<exists> a b. x = (a,b)" by auto
@@ -43,7 +43,7 @@
let ?a' = "a div ?g"
let ?b' = "b div ?g"
let ?g' = "gcd ?a' ?b'"
- from anz bnz have "?g \<noteq> 0" by simp with int_gcd_ge_0[of a b]
+ from anz bnz have "?g \<noteq> 0" by simp with gcd_ge_0_int[of a b]
have gpos: "?g > 0" by arith
have gdvd: "?g dvd a" "?g dvd b" by arith+
from zdvd_mult_div_cancel[OF gdvd(1)] zdvd_mult_div_cancel[OF gdvd(2)]
@@ -51,7 +51,7 @@
have nz':"?a' \<noteq> 0" "?b' \<noteq> 0"
by - (rule notI, simp)+
from anz bnz have stupid: "a \<noteq> 0 \<or> b \<noteq> 0" by arith
- from int_div_gcd_coprime[OF stupid] have gp1: "?g' = 1" .
+ from div_gcd_coprime_int[OF stupid] have gp1: "?g' = 1" .
from bnz have "b < 0 \<or> b > 0" by arith
moreover
{assume b: "b > 0"
@@ -137,7 +137,7 @@
lemma Ninv_normN[simp]: "isnormNum x \<Longrightarrow> isnormNum (Ninv x)"
by (simp add: Ninv_def isnormNum_def split_def)
- (cases "fst x = 0", auto simp add: int_gcd_commute)
+ (cases "fst x = 0", auto simp add: gcd_commute_int)
lemma isnormNum_int[simp]:
"isnormNum 0\<^sub>N" "isnormNum (1::int)\<^sub>N" "i \<noteq> 0 \<Longrightarrow> isnormNum i\<^sub>N"
@@ -203,7 +203,7 @@
from prems have eq:"a * b' = a'*b"
by (simp add: INum_def eq_divide_eq divide_eq_eq of_int_mult[symmetric] del: of_int_mult)
from prems have gcd1: "gcd a b = 1" "gcd b a = 1" "gcd a' b' = 1" "gcd b' a' = 1"
- by (simp_all add: isnormNum_def add: int_gcd_commute)
+ by (simp_all add: isnormNum_def add: gcd_commute_int)
from eq have raw_dvd: "a dvd a'*b" "b dvd b'*a" "a' dvd a*b'" "b' dvd b*a'"
apply -
apply algebra
@@ -211,8 +211,8 @@
apply simp
apply algebra
done
- from zdvd_dvd_eq[OF bz int_coprime_dvd_mult[OF gcd1(2) raw_dvd(2)]
- int_coprime_dvd_mult[OF gcd1(4) raw_dvd(4)]]
+ from zdvd_dvd_eq[OF bz coprime_dvd_mult_int[OF gcd1(2) raw_dvd(2)]
+ coprime_dvd_mult_int[OF gcd1(4) raw_dvd(4)]]
have eq1: "b = b'" using pos by arith
with eq have "a = a'" using pos by simp
with eq1 have ?rhs by simp}
@@ -297,8 +297,8 @@
let ?g = "gcd (a * b' + b * a') (b*b')"
have gz: "?g \<noteq> 0" using z by simp
have ?thesis using aa' bb' z gz
- of_int_div[where ?'a = 'a, OF gz int_gcd_dvd1[where x="a * b' + b * a'" and y="b*b'"]] of_int_div[where ?'a = 'a,
- OF gz int_gcd_dvd2[where x="a * b' + b * a'" and y="b*b'"]]
+ of_int_div[where ?'a = 'a, OF gz gcd_dvd1_int[where x="a * b' + b * a'" and y="b*b'"]] of_int_div[where ?'a = 'a,
+ OF gz gcd_dvd2_int[where x="a * b' + b * a'" and y="b*b'"]]
by (simp add: x y Nadd_def INum_def normNum_def Let_def add_divide_distrib)}
ultimately have ?thesis using aa' bb'
by (simp add: Nadd_def INum_def normNum_def x y Let_def) }
@@ -319,8 +319,8 @@
{assume z: "a \<noteq> 0" "a' \<noteq> 0" "b \<noteq> 0" "b' \<noteq> 0"
let ?g="gcd (a*a') (b*b')"
have gz: "?g \<noteq> 0" using z by simp
- from z of_int_div[where ?'a = 'a, OF gz int_gcd_dvd1[where x="a*a'" and y="b*b'"]]
- of_int_div[where ?'a = 'a , OF gz int_gcd_dvd2[where x="a*a'" and y="b*b'"]]
+ from z of_int_div[where ?'a = 'a, OF gz gcd_dvd1_int[where x="a*a'" and y="b*b'"]]
+ of_int_div[where ?'a = 'a , OF gz gcd_dvd2_int[where x="a*a'" and y="b*b'"]]
have ?thesis by (simp add: Nmul_def x y Let_def INum_def)}
ultimately show ?thesis by blast
qed
@@ -478,7 +478,7 @@
qed
lemma Nmul_commute: "isnormNum x \<Longrightarrow> isnormNum y \<Longrightarrow> x *\<^sub>N y = y *\<^sub>N x"
- by (simp add: Nmul_def split_def Let_def int_gcd_commute mult_commute)
+ by (simp add: Nmul_def split_def Let_def gcd_commute_int mult_commute)
lemma Nmul_assoc:
assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
--- a/src/HOL/Library/Efficient_Nat.thy Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/Library/Efficient_Nat.thy Thu Jul 09 08:55:42 2009 +0200
@@ -105,15 +105,9 @@
This can be accomplished by applying the following transformation rules:
*}
-lemma Suc_if_eq': "(\<And>n. f (Suc n) = h n) \<Longrightarrow> f 0 = g \<Longrightarrow>
- f n = (if n = 0 then g else h (n - 1))"
- by (cases n) simp_all
-
lemma Suc_if_eq: "(\<And>n. f (Suc n) \<equiv> h n) \<Longrightarrow> f 0 \<equiv> g \<Longrightarrow>
f n \<equiv> if n = 0 then g else h (n - 1)"
- by (rule eq_reflection, rule Suc_if_eq')
- (rule meta_eq_to_obj_eq, assumption,
- rule meta_eq_to_obj_eq, assumption)
+ by (rule eq_reflection) (cases n, simp_all)
lemma Suc_clause: "(\<And>n. P n (Suc n)) \<Longrightarrow> n \<noteq> 0 \<Longrightarrow> P (n - 1) n"
by (cases n) simp_all
@@ -129,14 +123,14 @@
setup {*
let
-fun gen_remove_suc Suc_if_eq dest_judgement thy thms =
+fun remove_suc thy thms =
let
val vname = Name.variant (map fst
(fold (Term.add_var_names o Thm.full_prop_of) thms [])) "n";
val cv = cterm_of thy (Var ((vname, 0), HOLogic.natT));
fun lhs_of th = snd (Thm.dest_comb
- (fst (Thm.dest_comb (dest_judgement (cprop_of th)))));
- fun rhs_of th = snd (Thm.dest_comb (dest_judgement (cprop_of th)));
+ (fst (Thm.dest_comb (cprop_of th))));
+ fun rhs_of th = snd (Thm.dest_comb (cprop_of th));
fun find_vars ct = (case term_of ct of
(Const (@{const_name Suc}, _) $ Var _) => [(cv, snd (Thm.dest_comb ct))]
| _ $ _ =>
@@ -156,7 +150,7 @@
(Drule.instantiate'
[SOME (ctyp_of_term ct)] [SOME (Thm.cabs cv ct),
SOME (Thm.cabs cv' (rhs_of th)), NONE, SOME cv']
- Suc_if_eq)) (Thm.forall_intr cv' th)
+ @{thm Suc_if_eq})) (Thm.forall_intr cv' th)
in
case map_filter (fn th'' =>
SOME (th'', singleton
@@ -169,21 +163,19 @@
end
in get_first mk_thms eqs end;
-fun gen_eqn_suc_preproc Suc_if_eq dest_judgement dest_lhs thy thms =
+fun eqn_suc_preproc thy thms =
let
- val dest = dest_lhs o prop_of;
+ val dest = fst o Logic.dest_equals o prop_of;
val contains_suc = exists_Const (fn (c, _) => c = @{const_name Suc});
in
if forall (can dest) thms andalso exists (contains_suc o dest) thms
- then perhaps_loop (gen_remove_suc Suc_if_eq dest_judgement thy) thms
+ then perhaps_loop (remove_suc thy) thms
else NONE
end;
-val eqn_suc_preproc = Code_Preproc.simple_functrans (gen_eqn_suc_preproc
- @{thm Suc_if_eq} I (fst o Logic.dest_equals));
+val eqn_suc_preproc1 = Code_Preproc.simple_functrans eqn_suc_preproc;
-fun eqn_suc_preproc' thy thms = gen_eqn_suc_preproc
- @{thm Suc_if_eq'} (snd o Thm.dest_comb) (fst o HOLogic.dest_eq o HOLogic.dest_Trueprop) thy thms
+fun eqn_suc_preproc2 thy thms = eqn_suc_preproc thy thms
|> the_default thms;
fun remove_suc_clause thy thms =
@@ -227,9 +219,9 @@
end;
in
- Codegen.add_preprocessor eqn_suc_preproc'
+ Codegen.add_preprocessor eqn_suc_preproc2
#> Codegen.add_preprocessor clause_suc_preproc
- #> Code_Preproc.add_functrans ("eqn_Suc", eqn_suc_preproc)
+ #> Code_Preproc.add_functrans ("eqn_Suc", eqn_suc_preproc1)
end;
*}
--- a/src/HOL/Library/Legacy_GCD.thy Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/Library/Legacy_GCD.thy Thu Jul 09 08:55:42 2009 +0200
@@ -409,7 +409,7 @@
{fix x y assume H: "a * x - b * y = d \<or> b * x - a * y = d"
have dv: "?g dvd a*x" "?g dvd b * y" "?g dvd b*x" "?g dvd a * y"
using dvd_mult2[OF gcd_dvd1[of a b]] dvd_mult2[OF gcd_dvd2[of a b]] by simp_all
- from nat_dvd_diff[OF dv(1,2)] nat_dvd_diff[OF dv(3,4)] H
+ from dvd_diff_nat[OF dv(1,2)] dvd_diff_nat[OF dv(3,4)] H
have ?rhs by auto}
ultimately show ?thesis by blast
qed
--- a/src/HOL/Library/Pocklington.thy Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/Library/Pocklington.thy Thu Jul 09 08:55:42 2009 +0200
@@ -846,7 +846,7 @@
from lh[unfolded nat_mod]
obtain q1 q2 where q12:"a ^ d + n * q1 = 1 + n * q2" by blast
hence "a ^ d + n * q1 - n * q2 = 1" by simp
- with nat_dvd_diff [OF dvd_add [OF divides_rexp[OF p(2), of d'] dvd_mult2[OF p(1), of q1]] dvd_mult2[OF p(1), of q2]] d' have "p dvd 1" by simp
+ with dvd_diff_nat [OF dvd_add [OF divides_rexp[OF p(2), of d'] dvd_mult2[OF p(1), of q1]] dvd_mult2[OF p(1), of q2]] d' have "p dvd 1" by simp
with p(3) have False by simp
hence ?rhs ..}
ultimately have ?rhs by blast}
--- a/src/HOL/Library/Primes.thy Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/Library/Primes.thy Thu Jul 09 08:55:42 2009 +0200
@@ -311,8 +311,8 @@
{fix e assume H: "e dvd a^n" "e dvd b^n"
from bezout_gcd_pow[of a n b] obtain x y
where xy: "a ^ n * x - b ^ n * y = ?gn \<or> b ^ n * x - a ^ n * y = ?gn" by blast
- from nat_dvd_diff [OF dvd_mult2[OF H(1), of x] dvd_mult2[OF H(2), of y]]
- nat_dvd_diff [OF dvd_mult2[OF H(2), of x] dvd_mult2[OF H(1), of y]] xy
+ from dvd_diff_nat [OF dvd_mult2[OF H(1), of x] dvd_mult2[OF H(2), of y]]
+ dvd_diff_nat [OF dvd_mult2[OF H(2), of x] dvd_mult2[OF H(1), of y]] xy
have "e dvd ?gn" by (cases "a ^ n * x - b ^ n * y = gcd a b ^ n", simp_all)}
hence th: "\<forall>e. e dvd a^n \<and> e dvd b^n \<longrightarrow> e dvd ?gn" by blast
from divides_exp[OF gcd_dvd1[of a b], of n] divides_exp[OF gcd_dvd2[of a b], of n] th
--- a/src/HOL/NewNumberTheory/Binomial.thy Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/NewNumberTheory/Binomial.thy Thu Jul 09 08:55:42 2009 +0200
@@ -109,68 +109,68 @@
subsection {* Factorial *}
-lemma nat_fact_zero [simp]: "fact (0::nat) = 1"
+lemma fact_zero_nat [simp]: "fact (0::nat) = 1"
by simp
-lemma int_fact_zero [simp]: "fact (0::int) = 1"
+lemma fact_zero_int [simp]: "fact (0::int) = 1"
by (simp add: fact_int_def)
-lemma nat_fact_one [simp]: "fact (1::nat) = 1"
+lemma fact_one_nat [simp]: "fact (1::nat) = 1"
by simp
-lemma nat_fact_Suc_0 [simp]: "fact (Suc 0) = Suc 0"
+lemma fact_Suc_0_nat [simp]: "fact (Suc 0) = Suc 0"
by (simp add: One_nat_def)
-lemma int_fact_one [simp]: "fact (1::int) = 1"
+lemma fact_one_int [simp]: "fact (1::int) = 1"
by (simp add: fact_int_def)
-lemma nat_fact_plus_one: "fact ((n::nat) + 1) = (n + 1) * fact n"
+lemma fact_plus_one_nat: "fact ((n::nat) + 1) = (n + 1) * fact n"
by simp
-lemma nat_fact_Suc: "fact (Suc n) = (Suc n) * fact n"
+lemma fact_Suc_nat: "fact (Suc n) = (Suc n) * fact n"
by (simp add: One_nat_def)
-lemma int_fact_plus_one:
+lemma fact_plus_one_int:
assumes "n >= 0"
shows "fact ((n::int) + 1) = (n + 1) * fact n"
- using prems by (rule nat_fact_plus_one [transferred])
+ using prems by (rule fact_plus_one_nat [transferred])
-lemma nat_fact_reduce: "(n::nat) > 0 \<Longrightarrow> fact n = n * fact (n - 1)"
+lemma fact_reduce_nat: "(n::nat) > 0 \<Longrightarrow> fact n = n * fact (n - 1)"
by simp
-lemma int_fact_reduce:
+lemma fact_reduce_int:
assumes "(n::int) > 0"
shows "fact n = n * fact (n - 1)"
using prems apply (subst tsub_eq [symmetric], auto)
- apply (rule nat_fact_reduce [transferred])
+ apply (rule fact_reduce_nat [transferred])
using prems apply auto
done
declare fact_nat.simps [simp del]
-lemma nat_fact_nonzero [simp]: "fact (n::nat) \<noteq> 0"
- apply (induct n rule: nat_induct')
- apply (auto simp add: nat_fact_plus_one)
+lemma fact_nonzero_nat [simp]: "fact (n::nat) \<noteq> 0"
+ apply (induct n rule: induct'_nat)
+ apply (auto simp add: fact_plus_one_nat)
done
-lemma int_fact_nonzero [simp]: "n >= 0 \<Longrightarrow> fact (n::int) ~= 0"
+lemma fact_nonzero_int [simp]: "n >= 0 \<Longrightarrow> fact (n::int) ~= 0"
by (simp add: fact_int_def)
-lemma nat_fact_gt_zero [simp]: "fact (n :: nat) > 0"
- by (insert nat_fact_nonzero [of n], arith)
+lemma fact_gt_zero_nat [simp]: "fact (n :: nat) > 0"
+ by (insert fact_nonzero_nat [of n], arith)
-lemma int_fact_gt_zero [simp]: "n >= 0 \<Longrightarrow> fact (n :: int) > 0"
+lemma fact_gt_zero_int [simp]: "n >= 0 \<Longrightarrow> fact (n :: int) > 0"
by (auto simp add: fact_int_def)
-lemma nat_fact_ge_one [simp]: "fact (n :: nat) >= 1"
- by (insert nat_fact_nonzero [of n], arith)
+lemma fact_ge_one_nat [simp]: "fact (n :: nat) >= 1"
+ by (insert fact_nonzero_nat [of n], arith)
-lemma nat_fact_ge_Suc_0 [simp]: "fact (n :: nat) >= Suc 0"
- by (insert nat_fact_nonzero [of n], arith)
+lemma fact_ge_Suc_0_nat [simp]: "fact (n :: nat) >= Suc 0"
+ by (insert fact_nonzero_nat [of n], arith)
-lemma int_fact_ge_one [simp]: "n >= 0 \<Longrightarrow> fact (n :: int) >= 1"
+lemma fact_ge_one_int [simp]: "n >= 0 \<Longrightarrow> fact (n :: int) >= 1"
apply (auto simp add: fact_int_def)
apply (subgoal_tac "1 = int 1")
apply (erule ssubst)
@@ -178,41 +178,41 @@
apply auto
done
-lemma nat_dvd_fact [rule_format]: "1 <= m \<longrightarrow> m <= n \<longrightarrow> m dvd fact (n::nat)"
- apply (induct n rule: nat_induct')
- apply (auto simp add: nat_fact_plus_one)
+lemma dvd_fact_nat [rule_format]: "1 <= m \<longrightarrow> m <= n \<longrightarrow> m dvd fact (n::nat)"
+ apply (induct n rule: induct'_nat)
+ apply (auto simp add: fact_plus_one_nat)
apply (subgoal_tac "m = n + 1")
apply auto
done
-lemma int_dvd_fact [rule_format]: "1 <= m \<longrightarrow> m <= n \<longrightarrow> m dvd fact (n::int)"
+lemma dvd_fact_int [rule_format]: "1 <= m \<longrightarrow> m <= n \<longrightarrow> m dvd fact (n::int)"
apply (case_tac "1 <= n")
apply (induct n rule: int_ge_induct)
- apply (auto simp add: int_fact_plus_one)
+ apply (auto simp add: fact_plus_one_int)
apply (subgoal_tac "m = i + 1")
apply auto
done
-lemma nat_interval_plus_one: "(i::nat) <= j + 1 \<Longrightarrow>
+lemma interval_plus_one_nat: "(i::nat) <= j + 1 \<Longrightarrow>
{i..j+1} = {i..j} Un {j+1}"
by auto
-lemma int_interval_plus_one: "(i::int) <= j + 1 \<Longrightarrow> {i..j+1} = {i..j} Un {j+1}"
+lemma interval_plus_one_int: "(i::int) <= j + 1 \<Longrightarrow> {i..j+1} = {i..j} Un {j+1}"
by auto
-lemma nat_fact_altdef: "fact (n::nat) = (PROD i:{1..n}. i)"
- apply (induct n rule: nat_induct')
+lemma fact_altdef_nat: "fact (n::nat) = (PROD i:{1..n}. i)"
+ apply (induct n rule: induct'_nat)
apply force
- apply (subst nat_fact_plus_one)
- apply (subst nat_interval_plus_one)
+ apply (subst fact_plus_one_nat)
+ apply (subst interval_plus_one_nat)
apply auto
done
-lemma int_fact_altdef: "n >= 0 \<Longrightarrow> fact (n::int) = (PROD i:{1..n}. i)"
+lemma fact_altdef_int: "n >= 0 \<Longrightarrow> fact (n::int) = (PROD i:{1..n}. i)"
apply (induct n rule: int_ge_induct)
apply force
- apply (subst int_fact_plus_one, assumption)
- apply (subst int_interval_plus_one)
+ apply (subst fact_plus_one_int, assumption)
+ apply (subst interval_plus_one_int)
apply auto
done
@@ -220,8 +220,8 @@
lemma next_prime_bound: "\<exists>(p::nat). prime p \<and> n < p \<and> p <= fact n + 1"
proof-
- have f1: "fact n + 1 \<noteq> 1" using nat_fact_ge_one [of n] by arith
- from nat_prime_factor [OF f1]
+ have f1: "fact n + 1 \<noteq> 1" using fact_ge_one_nat [of n] by arith
+ from prime_factor_nat [OF f1]
obtain p where "prime p" and "p dvd fact n + 1" by auto
hence "p \<le> fact n + 1"
by (intro dvd_imp_le, auto)
@@ -229,9 +229,9 @@
from `prime p` have "p \<ge> 1"
by (cases p, simp_all)
with `p <= n` have "p dvd fact n"
- by (intro nat_dvd_fact)
+ by (intro dvd_fact_nat)
with `p dvd fact n + 1` have "p dvd fact n + 1 - fact n"
- by (rule nat_dvd_diff)
+ by (rule dvd_diff_nat)
hence "p dvd 1" by simp
hence "p <= 1" by auto
moreover from `prime p` have "p > 1" by auto
@@ -256,95 +256,94 @@
subsection {* Binomial coefficients *}
-lemma nat_choose_zero [simp]: "(n::nat) choose 0 = 1"
+lemma choose_zero_nat [simp]: "(n::nat) choose 0 = 1"
by simp
-lemma int_choose_zero [simp]: "n \<ge> 0 \<Longrightarrow> (n::int) choose 0 = 1"
+lemma choose_zero_int [simp]: "n \<ge> 0 \<Longrightarrow> (n::int) choose 0 = 1"
by (simp add: binomial_int_def)
-lemma nat_zero_choose [rule_format,simp]: "ALL (k::nat) > n. n choose k = 0"
- by (induct n rule: nat_induct', auto)
+lemma zero_choose_nat [rule_format,simp]: "ALL (k::nat) > n. n choose k = 0"
+ by (induct n rule: induct'_nat, auto)
-lemma int_zero_choose [rule_format,simp]: "(k::int) > n \<Longrightarrow> n choose k = 0"
+lemma zero_choose_int [rule_format,simp]: "(k::int) > n \<Longrightarrow> n choose k = 0"
unfolding binomial_int_def apply (case_tac "n < 0")
apply force
apply (simp del: binomial_nat.simps)
done
-lemma nat_choose_reduce: "(n::nat) > 0 \<Longrightarrow> 0 < k \<Longrightarrow>
+lemma choose_reduce_nat: "(n::nat) > 0 \<Longrightarrow> 0 < k \<Longrightarrow>
(n choose k) = ((n - 1) choose k) + ((n - 1) choose (k - 1))"
by simp
-lemma int_choose_reduce: "(n::int) > 0 \<Longrightarrow> 0 < k \<Longrightarrow>
+lemma choose_reduce_int: "(n::int) > 0 \<Longrightarrow> 0 < k \<Longrightarrow>
(n choose k) = ((n - 1) choose k) + ((n - 1) choose (k - 1))"
- unfolding binomial_int_def apply (subst nat_choose_reduce)
+ unfolding binomial_int_def apply (subst choose_reduce_nat)
apply (auto simp del: binomial_nat.simps
simp add: nat_diff_distrib)
done
-lemma nat_choose_plus_one: "((n::nat) + 1) choose (k + 1) =
+lemma choose_plus_one_nat: "((n::nat) + 1) choose (k + 1) =
(n choose (k + 1)) + (n choose k)"
- by (simp add: nat_choose_reduce)
+ by (simp add: choose_reduce_nat)
-lemma nat_choose_Suc: "(Suc n) choose (Suc k) =
+lemma choose_Suc_nat: "(Suc n) choose (Suc k) =
(n choose (Suc k)) + (n choose k)"
- by (simp add: nat_choose_reduce One_nat_def)
+ by (simp add: choose_reduce_nat One_nat_def)
-lemma int_choose_plus_one: "n \<ge> 0 \<Longrightarrow> k \<ge> 0 \<Longrightarrow> ((n::int) + 1) choose (k + 1) =
+lemma choose_plus_one_int: "n \<ge> 0 \<Longrightarrow> k \<ge> 0 \<Longrightarrow> ((n::int) + 1) choose (k + 1) =
(n choose (k + 1)) + (n choose k)"
- by (simp add: binomial_int_def nat_choose_plus_one nat_add_distrib
- del: binomial_nat.simps)
+ by (simp add: binomial_int_def choose_plus_one_nat nat_add_distrib del: binomial_nat.simps)
declare binomial_nat.simps [simp del]
-lemma nat_choose_self [simp]: "((n::nat) choose n) = 1"
- by (induct n rule: nat_induct', auto simp add: nat_choose_plus_one)
+lemma choose_self_nat [simp]: "((n::nat) choose n) = 1"
+ by (induct n rule: induct'_nat, auto simp add: choose_plus_one_nat)
-lemma int_choose_self [simp]: "n \<ge> 0 \<Longrightarrow> ((n::int) choose n) = 1"
+lemma choose_self_int [simp]: "n \<ge> 0 \<Longrightarrow> ((n::int) choose n) = 1"
by (auto simp add: binomial_int_def)
-lemma nat_choose_one [simp]: "(n::nat) choose 1 = n"
- by (induct n rule: nat_induct', auto simp add: nat_choose_reduce)
+lemma choose_one_nat [simp]: "(n::nat) choose 1 = n"
+ by (induct n rule: induct'_nat, auto simp add: choose_reduce_nat)
-lemma int_choose_one [simp]: "n \<ge> 0 \<Longrightarrow> (n::int) choose 1 = n"
+lemma choose_one_int [simp]: "n \<ge> 0 \<Longrightarrow> (n::int) choose 1 = n"
by (auto simp add: binomial_int_def)
-lemma nat_plus_one_choose_self [simp]: "(n::nat) + 1 choose n = n + 1"
- apply (induct n rule: nat_induct', force)
+lemma plus_one_choose_self_nat [simp]: "(n::nat) + 1 choose n = n + 1"
+ apply (induct n rule: induct'_nat, force)
apply (case_tac "n = 0")
apply auto
- apply (subst nat_choose_reduce)
+ apply (subst choose_reduce_nat)
apply (auto simp add: One_nat_def)
(* natdiff_cancel_numerals introduces Suc *)
done
-lemma nat_Suc_choose_self [simp]: "(Suc n) choose n = Suc n"
- using nat_plus_one_choose_self by (simp add: One_nat_def)
+lemma Suc_choose_self_nat [simp]: "(Suc n) choose n = Suc n"
+ using plus_one_choose_self_nat by (simp add: One_nat_def)
-lemma int_plus_one_choose_self [rule_format, simp]:
+lemma plus_one_choose_self_int [rule_format, simp]:
"(n::int) \<ge> 0 \<longrightarrow> n + 1 choose n = n + 1"
by (auto simp add: binomial_int_def nat_add_distrib)
(* bounded quantification doesn't work with the unicode characters? *)
-lemma nat_choose_pos [rule_format]: "ALL k <= (n::nat).
+lemma choose_pos_nat [rule_format]: "ALL k <= (n::nat).
((n::nat) choose k) > 0"
- apply (induct n rule: nat_induct')
+ apply (induct n rule: induct'_nat)
apply force
apply clarify
apply (case_tac "k = 0")
apply force
- apply (subst nat_choose_reduce)
+ apply (subst choose_reduce_nat)
apply auto
done
-lemma int_choose_pos: "n \<ge> 0 \<Longrightarrow> k >= 0 \<Longrightarrow> k \<le> n \<Longrightarrow>
+lemma choose_pos_int: "n \<ge> 0 \<Longrightarrow> k >= 0 \<Longrightarrow> k \<le> n \<Longrightarrow>
((n::int) choose k) > 0"
- by (auto simp add: binomial_int_def nat_choose_pos)
+ by (auto simp add: binomial_int_def choose_pos_nat)
lemma binomial_induct [rule_format]: "(ALL (n::nat). P n n) \<longrightarrow>
(ALL n. P (n + 1) 0) \<longrightarrow> (ALL n. (ALL k < n. P n k \<longrightarrow> P n (k + 1) \<longrightarrow>
P (n + 1) (k + 1))) \<longrightarrow> (ALL k <= n. P n k)"
- apply (induct n rule: nat_induct')
+ apply (induct n rule: induct'_nat)
apply auto
apply (case_tac "k = 0")
apply auto
@@ -355,7 +354,7 @@
apply auto
done
-lemma nat_choose_altdef_aux: "(k::nat) \<le> n \<Longrightarrow>
+lemma choose_altdef_aux_nat: "(k::nat) \<le> n \<Longrightarrow>
fact k * fact (n - k) * (n choose k) = fact n"
apply (rule binomial_induct [of _ k n])
apply auto
@@ -364,22 +363,22 @@
assume less: "k < n"
assume ih1: "fact k * fact (n - k) * (n choose k) = fact n"
hence one: "fact (k + 1) * fact (n - k) * (n choose k) = (k + 1) * fact n"
- by (subst nat_fact_plus_one, auto)
+ by (subst fact_plus_one_nat, auto)
assume ih2: "fact (k + 1) * fact (n - (k + 1)) * (n choose (k + 1)) =
fact n"
with less have "fact (k + 1) * fact ((n - (k + 1)) + 1) *
(n choose (k + 1)) = (n - k) * fact n"
- by (subst (2) nat_fact_plus_one, auto)
+ by (subst (2) fact_plus_one_nat, auto)
with less have two: "fact (k + 1) * fact (n - k) * (n choose (k + 1)) =
(n - k) * fact n" by simp
have "fact (k + 1) * fact (n - k) * (n + 1 choose (k + 1)) =
fact (k + 1) * fact (n - k) * (n choose (k + 1)) +
fact (k + 1) * fact (n - k) * (n choose k)"
- by (subst nat_choose_reduce, auto simp add: ring_simps)
+ by (subst choose_reduce_nat, auto simp add: ring_simps)
also note one
also note two
also with less have "(n - k) * fact n + (k + 1) * fact n= fact (n + 1)"
- apply (subst nat_fact_plus_one)
+ apply (subst fact_plus_one_nat)
apply (subst left_distrib [symmetric])
apply simp
done
@@ -387,43 +386,43 @@
fact (n + 1)" .
qed
-lemma nat_choose_altdef: "(k::nat) \<le> n \<Longrightarrow>
+lemma choose_altdef_nat: "(k::nat) \<le> n \<Longrightarrow>
n choose k = fact n div (fact k * fact (n - k))"
- apply (frule nat_choose_altdef_aux)
+ apply (frule choose_altdef_aux_nat)
apply (erule subst)
apply (simp add: mult_ac)
done
-lemma int_choose_altdef:
+lemma choose_altdef_int:
assumes "(0::int) <= k" and "k <= n"
shows "n choose k = fact n div (fact k * fact (n - k))"
apply (subst tsub_eq [symmetric], rule prems)
- apply (rule nat_choose_altdef [transferred])
+ apply (rule choose_altdef_nat [transferred])
using prems apply auto
done
-lemma nat_choose_dvd: "(k::nat) \<le> n \<Longrightarrow> fact k * fact (n - k) dvd fact n"
- unfolding dvd_def apply (frule nat_choose_altdef_aux)
+lemma choose_dvd_nat: "(k::nat) \<le> n \<Longrightarrow> fact k * fact (n - k) dvd fact n"
+ unfolding dvd_def apply (frule choose_altdef_aux_nat)
(* why don't blast and auto get this??? *)
apply (rule exI)
apply (erule sym)
done
-lemma int_choose_dvd:
+lemma choose_dvd_int:
assumes "(0::int) <= k" and "k <= n"
shows "fact k * fact (n - k) dvd fact n"
apply (subst tsub_eq [symmetric], rule prems)
- apply (rule nat_choose_dvd [transferred])
+ apply (rule choose_dvd_nat [transferred])
using prems apply auto
done
(* generalizes Tobias Nipkow's proof to any commutative semiring *)
theorem binomial: "(a+b::'a::{comm_ring_1,power})^n =
(SUM k=0..n. (of_nat (n choose k)) * a^k * b^(n-k))" (is "?P n")
-proof (induct n rule: nat_induct')
+proof (induct n rule: induct'_nat)
show "?P 0" by simp
next
fix n
@@ -455,7 +454,7 @@
"... = a^(n+1) + b^(n+1) +
(SUM k=1..n. of_nat(n+1 choose k) * a^k * b^(n+1-k))"
by (auto simp add: ring_simps setsum_addf [symmetric]
- nat_choose_reduce)
+ choose_reduce_nat)
also have "... = (SUM k=0..n+1. of_nat (n+1 choose k) * a^k * b^(n+1-k))"
using decomp by (simp add: ring_simps)
finally show "?P (n + 1)" by simp
@@ -464,7 +463,7 @@
lemma set_explicit: "{S. S = T \<and> P S} = (if P T then {T} else {})"
by auto
-lemma nat_card_subsets [rule_format]:
+lemma card_subsets_nat [rule_format]:
fixes S :: "'a set"
assumes "finite S"
shows "ALL k. card {T. T \<le> S \<and> card T = k} = card S choose k"
@@ -480,7 +479,7 @@
fix k
show "card {T. T \<subseteq> (insert x F) \<and> card T = k} =
card (insert x F) choose k" (is "?Q k")
- proof (induct k rule: nat_induct')
+ proof (induct k rule: induct'_nat)
from iassms have "{T. T \<le> (insert x F) \<and> card T = 0} = {{}}"
apply auto
apply (subst (asm) card_0_eq)
@@ -530,7 +529,7 @@
by auto
finally have "card ({T. T \<le> insert x F \<and> card T = k + 1}) =
card F choose (k + 1) + (card F choose k)".
- with iassms nat_choose_plus_one show ?thesis
+ with iassms choose_plus_one_nat show ?thesis
by auto
qed
qed
--- a/src/HOL/NewNumberTheory/Cong.thy Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/NewNumberTheory/Cong.thy Thu Jul 09 08:55:42 2009 +0200
@@ -35,18 +35,18 @@
subsection {* Turn off One_nat_def *}
-lemma nat_induct' [case_names zero plus1, induct type: nat]:
+lemma induct'_nat [case_names zero plus1, induct type: nat]:
"\<lbrakk> P (0::nat); !!n. P n \<Longrightarrow> P (n + 1)\<rbrakk> \<Longrightarrow> P n"
by (erule nat_induct) (simp add:One_nat_def)
-lemma nat_cases [case_names zero plus1, cases type: nat]:
+lemma cases_nat [case_names zero plus1, cases type: nat]:
"P (0::nat) \<Longrightarrow> (!!n. P (n + 1)) \<Longrightarrow> P n"
-by(metis nat_induct')
+by(metis induct'_nat)
lemma power_plus_one [simp]: "(x::'a::power)^(n + 1) = x * x^n"
by (simp add: One_nat_def)
-lemma nat_power_eq_one_eq [simp]:
+lemma power_eq_one_eq_nat [simp]:
"((x::nat)^m = 1) = (m = 0 | x = 1)"
by (induct m, auto)
@@ -147,55 +147,55 @@
subsection {* Congruence *}
(* was zcong_0, etc. *)
-lemma nat_cong_0 [simp, presburger]: "([(a::nat) = b] (mod 0)) = (a = b)"
+lemma cong_0_nat [simp, presburger]: "([(a::nat) = b] (mod 0)) = (a = b)"
by (unfold cong_nat_def, auto)
-lemma int_cong_0 [simp, presburger]: "([(a::int) = b] (mod 0)) = (a = b)"
+lemma cong_0_int [simp, presburger]: "([(a::int) = b] (mod 0)) = (a = b)"
by (unfold cong_int_def, auto)
-lemma nat_cong_1 [simp, presburger]: "[(a::nat) = b] (mod 1)"
+lemma cong_1_nat [simp, presburger]: "[(a::nat) = b] (mod 1)"
by (unfold cong_nat_def, auto)
-lemma nat_cong_Suc_0 [simp, presburger]: "[(a::nat) = b] (mod Suc 0)"
+lemma cong_Suc_0_nat [simp, presburger]: "[(a::nat) = b] (mod Suc 0)"
by (unfold cong_nat_def, auto simp add: One_nat_def)
-lemma int_cong_1 [simp, presburger]: "[(a::int) = b] (mod 1)"
+lemma cong_1_int [simp, presburger]: "[(a::int) = b] (mod 1)"
by (unfold cong_int_def, auto)
-lemma nat_cong_refl [simp]: "[(k::nat) = k] (mod m)"
+lemma cong_refl_nat [simp]: "[(k::nat) = k] (mod m)"
by (unfold cong_nat_def, auto)
-lemma int_cong_refl [simp]: "[(k::int) = k] (mod m)"
+lemma cong_refl_int [simp]: "[(k::int) = k] (mod m)"
by (unfold cong_int_def, auto)
-lemma nat_cong_sym: "[(a::nat) = b] (mod m) \<Longrightarrow> [b = a] (mod m)"
+lemma cong_sym_nat: "[(a::nat) = b] (mod m) \<Longrightarrow> [b = a] (mod m)"
by (unfold cong_nat_def, auto)
-lemma int_cong_sym: "[(a::int) = b] (mod m) \<Longrightarrow> [b = a] (mod m)"
+lemma cong_sym_int: "[(a::int) = b] (mod m) \<Longrightarrow> [b = a] (mod m)"
by (unfold cong_int_def, auto)
-lemma nat_cong_sym_eq: "[(a::nat) = b] (mod m) = [b = a] (mod m)"
+lemma cong_sym_eq_nat: "[(a::nat) = b] (mod m) = [b = a] (mod m)"
by (unfold cong_nat_def, auto)
-lemma int_cong_sym_eq: "[(a::int) = b] (mod m) = [b = a] (mod m)"
+lemma cong_sym_eq_int: "[(a::int) = b] (mod m) = [b = a] (mod m)"
by (unfold cong_int_def, auto)
-lemma nat_cong_trans [trans]:
+lemma cong_trans_nat [trans]:
"[(a::nat) = b] (mod m) \<Longrightarrow> [b = c] (mod m) \<Longrightarrow> [a = c] (mod m)"
by (unfold cong_nat_def, auto)
-lemma int_cong_trans [trans]:
+lemma cong_trans_int [trans]:
"[(a::int) = b] (mod m) \<Longrightarrow> [b = c] (mod m) \<Longrightarrow> [a = c] (mod m)"
by (unfold cong_int_def, auto)
-lemma nat_cong_add:
+lemma cong_add_nat:
"[(a::nat) = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a + c = b + d] (mod m)"
apply (unfold cong_nat_def)
apply (subst (1 2) mod_add_eq)
apply simp
done
-lemma int_cong_add:
+lemma cong_add_int:
"[(a::int) = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a + c = b + d] (mod m)"
apply (unfold cong_int_def)
apply (subst (1 2) mod_add_left_eq)
@@ -203,35 +203,35 @@
apply simp
done
-lemma int_cong_diff:
+lemma cong_diff_int:
"[(a::int) = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a - c = b - d] (mod m)"
apply (unfold cong_int_def)
apply (subst (1 2) mod_diff_eq)
apply simp
done
-lemma int_cong_diff_aux:
+lemma cong_diff_aux_int:
"(a::int) >= c \<Longrightarrow> b >= d \<Longrightarrow> [(a::int) = b] (mod m) \<Longrightarrow>
[c = d] (mod m) \<Longrightarrow> [tsub a c = tsub b d] (mod m)"
apply (subst (1 2) tsub_eq)
- apply (auto intro: int_cong_diff)
+ apply (auto intro: cong_diff_int)
done;
-lemma nat_cong_diff:
+lemma cong_diff_nat:
assumes "(a::nat) >= c" and "b >= d" and "[a = b] (mod m)" and
"[c = d] (mod m)"
shows "[a - c = b - d] (mod m)"
- using prems by (rule int_cong_diff_aux [transferred]);
+ using prems by (rule cong_diff_aux_int [transferred]);
-lemma nat_cong_mult:
+lemma cong_mult_nat:
"[(a::nat) = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a * c = b * d] (mod m)"
apply (unfold cong_nat_def)
apply (subst (1 2) mod_mult_eq)
apply simp
done
-lemma int_cong_mult:
+lemma cong_mult_int:
"[(a::int) = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a * c = b * d] (mod m)"
apply (unfold cong_int_def)
apply (subst (1 2) zmod_zmult1_eq)
@@ -240,197 +240,197 @@
apply simp
done
-lemma nat_cong_exp: "[(x::nat) = y] (mod n) \<Longrightarrow> [x^k = y^k] (mod n)"
+lemma cong_exp_nat: "[(x::nat) = y] (mod n) \<Longrightarrow> [x^k = y^k] (mod n)"
apply (induct k)
- apply (auto simp add: nat_cong_refl nat_cong_mult)
+ apply (auto simp add: cong_refl_nat cong_mult_nat)
done
-lemma int_cong_exp: "[(x::int) = y] (mod n) \<Longrightarrow> [x^k = y^k] (mod n)"
+lemma cong_exp_int: "[(x::int) = y] (mod n) \<Longrightarrow> [x^k = y^k] (mod n)"
apply (induct k)
- apply (auto simp add: int_cong_refl int_cong_mult)
+ apply (auto simp add: cong_refl_int cong_mult_int)
done
-lemma nat_cong_setsum [rule_format]:
+lemma cong_setsum_nat [rule_format]:
"(ALL x: A. [((f x)::nat) = g x] (mod m)) \<longrightarrow>
[(SUM x:A. f x) = (SUM x:A. g x)] (mod m)"
apply (case_tac "finite A")
apply (induct set: finite)
- apply (auto intro: nat_cong_add)
+ apply (auto intro: cong_add_nat)
done
-lemma int_cong_setsum [rule_format]:
+lemma cong_setsum_int [rule_format]:
"(ALL x: A. [((f x)::int) = g x] (mod m)) \<longrightarrow>
[(SUM x:A. f x) = (SUM x:A. g x)] (mod m)"
apply (case_tac "finite A")
apply (induct set: finite)
- apply (auto intro: int_cong_add)
+ apply (auto intro: cong_add_int)
done
-lemma nat_cong_setprod [rule_format]:
+lemma cong_setprod_nat [rule_format]:
"(ALL x: A. [((f x)::nat) = g x] (mod m)) \<longrightarrow>
[(PROD x:A. f x) = (PROD x:A. g x)] (mod m)"
apply (case_tac "finite A")
apply (induct set: finite)
- apply (auto intro: nat_cong_mult)
+ apply (auto intro: cong_mult_nat)
done
-lemma int_cong_setprod [rule_format]:
+lemma cong_setprod_int [rule_format]:
"(ALL x: A. [((f x)::int) = g x] (mod m)) \<longrightarrow>
[(PROD x:A. f x) = (PROD x:A. g x)] (mod m)"
apply (case_tac "finite A")
apply (induct set: finite)
- apply (auto intro: int_cong_mult)
+ apply (auto intro: cong_mult_int)
done
-lemma nat_cong_scalar: "[(a::nat)= b] (mod m) \<Longrightarrow> [a * k = b * k] (mod m)"
- by (rule nat_cong_mult, simp_all)
+lemma cong_scalar_nat: "[(a::nat)= b] (mod m) \<Longrightarrow> [a * k = b * k] (mod m)"
+ by (rule cong_mult_nat, simp_all)
-lemma int_cong_scalar: "[(a::int)= b] (mod m) \<Longrightarrow> [a * k = b * k] (mod m)"
- by (rule int_cong_mult, simp_all)
+lemma cong_scalar_int: "[(a::int)= b] (mod m) \<Longrightarrow> [a * k = b * k] (mod m)"
+ by (rule cong_mult_int, simp_all)
-lemma nat_cong_scalar2: "[(a::nat)= b] (mod m) \<Longrightarrow> [k * a = k * b] (mod m)"
- by (rule nat_cong_mult, simp_all)
+lemma cong_scalar2_nat: "[(a::nat)= b] (mod m) \<Longrightarrow> [k * a = k * b] (mod m)"
+ by (rule cong_mult_nat, simp_all)
-lemma int_cong_scalar2: "[(a::int)= b] (mod m) \<Longrightarrow> [k * a = k * b] (mod m)"
- by (rule int_cong_mult, simp_all)
+lemma cong_scalar2_int: "[(a::int)= b] (mod m) \<Longrightarrow> [k * a = k * b] (mod m)"
+ by (rule cong_mult_int, simp_all)
-lemma nat_cong_mult_self: "[(a::nat) * m = 0] (mod m)"
+lemma cong_mult_self_nat: "[(a::nat) * m = 0] (mod m)"
by (unfold cong_nat_def, auto)
-lemma int_cong_mult_self: "[(a::int) * m = 0] (mod m)"
+lemma cong_mult_self_int: "[(a::int) * m = 0] (mod m)"
by (unfold cong_int_def, auto)
-lemma int_cong_eq_diff_cong_0: "[(a::int) = b] (mod m) = [a - b = 0] (mod m)"
+lemma cong_eq_diff_cong_0_int: "[(a::int) = b] (mod m) = [a - b = 0] (mod m)"
apply (rule iffI)
- apply (erule int_cong_diff [of a b m b b, simplified])
- apply (erule int_cong_add [of "a - b" 0 m b b, simplified])
+ apply (erule cong_diff_int [of a b m b b, simplified])
+ apply (erule cong_add_int [of "a - b" 0 m b b, simplified])
done
-lemma int_cong_eq_diff_cong_0_aux: "a >= b \<Longrightarrow>
+lemma cong_eq_diff_cong_0_aux_int: "a >= b \<Longrightarrow>
[(a::int) = b] (mod m) = [tsub a b = 0] (mod m)"
- by (subst tsub_eq, assumption, rule int_cong_eq_diff_cong_0)
+ by (subst tsub_eq, assumption, rule cong_eq_diff_cong_0_int)
-lemma nat_cong_eq_diff_cong_0:
+lemma cong_eq_diff_cong_0_nat:
assumes "(a::nat) >= b"
shows "[a = b] (mod m) = [a - b = 0] (mod m)"
- using prems by (rule int_cong_eq_diff_cong_0_aux [transferred])
+ using prems by (rule cong_eq_diff_cong_0_aux_int [transferred])
-lemma nat_cong_diff_cong_0':
+lemma cong_diff_cong_0'_nat:
"[(x::nat) = y] (mod n) \<longleftrightarrow>
(if x <= y then [y - x = 0] (mod n) else [x - y = 0] (mod n))"
apply (case_tac "y <= x")
- apply (frule nat_cong_eq_diff_cong_0 [where m = n])
+ apply (frule cong_eq_diff_cong_0_nat [where m = n])
apply auto [1]
apply (subgoal_tac "x <= y")
- apply (frule nat_cong_eq_diff_cong_0 [where m = n])
- apply (subst nat_cong_sym_eq)
+ apply (frule cong_eq_diff_cong_0_nat [where m = n])
+ apply (subst cong_sym_eq_nat)
apply auto
done
-lemma nat_cong_altdef: "(a::nat) >= b \<Longrightarrow> [a = b] (mod m) = (m dvd (a - b))"
- apply (subst nat_cong_eq_diff_cong_0, assumption)
+lemma cong_altdef_nat: "(a::nat) >= b \<Longrightarrow> [a = b] (mod m) = (m dvd (a - b))"
+ apply (subst cong_eq_diff_cong_0_nat, assumption)
apply (unfold cong_nat_def)
apply (simp add: dvd_eq_mod_eq_0 [symmetric])
done
-lemma int_cong_altdef: "[(a::int) = b] (mod m) = (m dvd (a - b))"
- apply (subst int_cong_eq_diff_cong_0)
+lemma cong_altdef_int: "[(a::int) = b] (mod m) = (m dvd (a - b))"
+ apply (subst cong_eq_diff_cong_0_int)
apply (unfold cong_int_def)
apply (simp add: dvd_eq_mod_eq_0 [symmetric])
done
-lemma int_cong_abs: "[(x::int) = y] (mod abs m) = [x = y] (mod m)"
- by (simp add: int_cong_altdef)
+lemma cong_abs_int: "[(x::int) = y] (mod abs m) = [x = y] (mod m)"
+ by (simp add: cong_altdef_int)
-lemma int_cong_square:
+lemma cong_square_int:
"\<lbrakk> prime (p::int); 0 < a; [a * a = 1] (mod p) \<rbrakk>
\<Longrightarrow> [a = 1] (mod p) \<or> [a = - 1] (mod p)"
- apply (simp only: int_cong_altdef)
- apply (subst int_prime_dvd_mult_eq [symmetric], assumption)
+ apply (simp only: cong_altdef_int)
+ apply (subst prime_dvd_mult_eq_int [symmetric], assumption)
(* any way around this? *)
apply (subgoal_tac "a * a - 1 = (a - 1) * (a - -1)")
apply (auto simp add: ring_simps)
done
-lemma int_cong_mult_rcancel:
+lemma cong_mult_rcancel_int:
"coprime k (m::int) \<Longrightarrow> [a * k = b * k] (mod m) = [a = b] (mod m)"
- apply (subst (1 2) int_cong_altdef)
+ apply (subst (1 2) cong_altdef_int)
apply (subst left_diff_distrib [symmetric])
- apply (rule int_coprime_dvd_mult_iff)
- apply (subst int_gcd_commute, assumption)
+ apply (rule coprime_dvd_mult_iff_int)
+ apply (subst gcd_commute_int, assumption)
done
-lemma nat_cong_mult_rcancel:
+lemma cong_mult_rcancel_nat:
assumes "coprime k (m::nat)"
shows "[a * k = b * k] (mod m) = [a = b] (mod m)"
- apply (rule int_cong_mult_rcancel [transferred])
+ apply (rule cong_mult_rcancel_int [transferred])
using prems apply auto
done
-lemma nat_cong_mult_lcancel:
+lemma cong_mult_lcancel_nat:
"coprime k (m::nat) \<Longrightarrow> [k * a = k * b ] (mod m) = [a = b] (mod m)"
- by (simp add: mult_commute nat_cong_mult_rcancel)
+ by (simp add: mult_commute cong_mult_rcancel_nat)
-lemma int_cong_mult_lcancel:
+lemma cong_mult_lcancel_int:
"coprime k (m::int) \<Longrightarrow> [k * a = k * b] (mod m) = [a = b] (mod m)"
- by (simp add: mult_commute int_cong_mult_rcancel)
+ by (simp add: mult_commute cong_mult_rcancel_int)
(* was zcong_zgcd_zmult_zmod *)
-lemma int_coprime_cong_mult:
+lemma coprime_cong_mult_int:
"[(a::int) = b] (mod m) \<Longrightarrow> [a = b] (mod n) \<Longrightarrow> coprime m n
\<Longrightarrow> [a = b] (mod m * n)"
- apply (simp only: int_cong_altdef)
- apply (erule (2) int_divides_mult)
+ apply (simp only: cong_altdef_int)
+ apply (erule (2) divides_mult_int)
done
-lemma nat_coprime_cong_mult:
+lemma coprime_cong_mult_nat:
assumes "[(a::nat) = b] (mod m)" and "[a = b] (mod n)" and "coprime m n"
shows "[a = b] (mod m * n)"
- apply (rule int_coprime_cong_mult [transferred])
+ apply (rule coprime_cong_mult_int [transferred])
using prems apply auto
done
-lemma nat_cong_less_imp_eq: "0 \<le> (a::nat) \<Longrightarrow>
+lemma cong_less_imp_eq_nat: "0 \<le> (a::nat) \<Longrightarrow>
a < m \<Longrightarrow> 0 \<le> b \<Longrightarrow> b < m \<Longrightarrow> [a = b] (mod m) \<Longrightarrow> a = b"
by (auto simp add: cong_nat_def mod_pos_pos_trivial)
-lemma int_cong_less_imp_eq: "0 \<le> (a::int) \<Longrightarrow>
+lemma cong_less_imp_eq_int: "0 \<le> (a::int) \<Longrightarrow>
a < m \<Longrightarrow> 0 \<le> b \<Longrightarrow> b < m \<Longrightarrow> [a = b] (mod m) \<Longrightarrow> a = b"
by (auto simp add: cong_int_def mod_pos_pos_trivial)
-lemma nat_cong_less_unique:
+lemma cong_less_unique_nat:
"0 < (m::nat) \<Longrightarrow> (\<exists>!b. 0 \<le> b \<and> b < m \<and> [a = b] (mod m))"
apply auto
apply (rule_tac x = "a mod m" in exI)
apply (unfold cong_nat_def, auto)
done
-lemma int_cong_less_unique:
+lemma cong_less_unique_int:
"0 < (m::int) \<Longrightarrow> (\<exists>!b. 0 \<le> b \<and> b < m \<and> [a = b] (mod m))"
apply auto
apply (rule_tac x = "a mod m" in exI)
apply (unfold cong_int_def, auto simp add: mod_pos_pos_trivial)
done
-lemma int_cong_iff_lin: "([(a::int) = b] (mod m)) = (\<exists>k. b = a + m * k)"
- apply (auto simp add: int_cong_altdef dvd_def ring_simps)
+lemma cong_iff_lin_int: "([(a::int) = b] (mod m)) = (\<exists>k. b = a + m * k)"
+ apply (auto simp add: cong_altdef_int dvd_def ring_simps)
apply (rule_tac [!] x = "-k" in exI, auto)
done
-lemma nat_cong_iff_lin: "([(a::nat) = b] (mod m)) =
+lemma cong_iff_lin_nat: "([(a::nat) = b] (mod m)) =
(\<exists>k1 k2. b + k1 * m = a + k2 * m)"
apply (rule iffI)
apply (case_tac "b <= a")
- apply (subst (asm) nat_cong_altdef, assumption)
+ apply (subst (asm) cong_altdef_nat, assumption)
apply (unfold dvd_def, auto)
apply (rule_tac x = k in exI)
apply (rule_tac x = 0 in exI)
apply (auto simp add: ring_simps)
- apply (subst (asm) nat_cong_sym_eq)
- apply (subst (asm) nat_cong_altdef)
+ apply (subst (asm) cong_sym_eq_nat)
+ apply (subst (asm) cong_altdef_nat)
apply force
apply (unfold dvd_def, auto)
apply (rule_tac x = 0 in exI)
@@ -443,50 +443,50 @@
apply auto
done
-lemma int_cong_gcd_eq: "[(a::int) = b] (mod m) \<Longrightarrow> gcd a m = gcd b m"
- apply (subst (asm) int_cong_iff_lin, auto)
+lemma cong_gcd_eq_int: "[(a::int) = b] (mod m) \<Longrightarrow> gcd a m = gcd b m"
+ apply (subst (asm) cong_iff_lin_int, auto)
apply (subst add_commute)
- apply (subst (2) int_gcd_commute)
+ apply (subst (2) gcd_commute_int)
apply (subst mult_commute)
- apply (subst int_gcd_add_mult)
- apply (rule int_gcd_commute)
+ apply (subst gcd_add_mult_int)
+ apply (rule gcd_commute_int)
done
-lemma nat_cong_gcd_eq:
+lemma cong_gcd_eq_nat:
assumes "[(a::nat) = b] (mod m)"
shows "gcd a m = gcd b m"
- apply (rule int_cong_gcd_eq [transferred])
+ apply (rule cong_gcd_eq_int [transferred])
using prems apply auto
done
-lemma nat_cong_imp_coprime: "[(a::nat) = b] (mod m) \<Longrightarrow> coprime a m \<Longrightarrow>
+lemma cong_imp_coprime_nat: "[(a::nat) = b] (mod m) \<Longrightarrow> coprime a m \<Longrightarrow>
coprime b m"
- by (auto simp add: nat_cong_gcd_eq)
+ by (auto simp add: cong_gcd_eq_nat)
-lemma int_cong_imp_coprime: "[(a::int) = b] (mod m) \<Longrightarrow> coprime a m \<Longrightarrow>
+lemma cong_imp_coprime_int: "[(a::int) = b] (mod m) \<Longrightarrow> coprime a m \<Longrightarrow>
coprime b m"
- by (auto simp add: int_cong_gcd_eq)
+ by (auto simp add: cong_gcd_eq_int)
-lemma nat_cong_cong_mod: "[(a::nat) = b] (mod m) =
+lemma cong_cong_mod_nat: "[(a::nat) = b] (mod m) =
[a mod m = b mod m] (mod m)"
by (auto simp add: cong_nat_def)
-lemma int_cong_cong_mod: "[(a::int) = b] (mod m) =
+lemma cong_cong_mod_int: "[(a::int) = b] (mod m) =
[a mod m = b mod m] (mod m)"
by (auto simp add: cong_int_def)
-lemma int_cong_minus [iff]: "[(a::int) = b] (mod -m) = [a = b] (mod m)"
- by (subst (1 2) int_cong_altdef, auto)
+lemma cong_minus_int [iff]: "[(a::int) = b] (mod -m) = [a = b] (mod m)"
+ by (subst (1 2) cong_altdef_int, auto)
-lemma nat_cong_zero [iff]: "[(a::nat) = b] (mod 0) = (a = b)"
+lemma cong_zero_nat [iff]: "[(a::nat) = b] (mod 0) = (a = b)"
by (auto simp add: cong_nat_def)
-lemma int_cong_zero [iff]: "[(a::int) = b] (mod 0) = (a = b)"
+lemma cong_zero_int [iff]: "[(a::int) = b] (mod 0) = (a = b)"
by (auto simp add: cong_int_def)
(*
-lemma int_mod_dvd_mod:
+lemma mod_dvd_mod_int:
"0 < (m::int) \<Longrightarrow> m dvd b \<Longrightarrow> (a mod b mod m) = (a mod m)"
apply (unfold dvd_def, auto)
apply (rule mod_mod_cancel)
@@ -497,79 +497,79 @@
assumes "0 < (m::nat)" and "m dvd b"
shows "(a mod b mod m) = (a mod m)"
- apply (rule int_mod_dvd_mod [transferred])
+ apply (rule mod_dvd_mod_int [transferred])
using prems apply auto
done
*)
-lemma nat_cong_add_lcancel:
+lemma cong_add_lcancel_nat:
"[(a::nat) + x = a + y] (mod n) \<longleftrightarrow> [x = y] (mod n)"
- by (simp add: nat_cong_iff_lin)
+ by (simp add: cong_iff_lin_nat)
-lemma int_cong_add_lcancel:
+lemma cong_add_lcancel_int:
"[(a::int) + x = a + y] (mod n) \<longleftrightarrow> [x = y] (mod n)"
- by (simp add: int_cong_iff_lin)
+ by (simp add: cong_iff_lin_int)
-lemma nat_cong_add_rcancel: "[(x::nat) + a = y + a] (mod n) \<longleftrightarrow> [x = y] (mod n)"
- by (simp add: nat_cong_iff_lin)
+lemma cong_add_rcancel_nat: "[(x::nat) + a = y + a] (mod n) \<longleftrightarrow> [x = y] (mod n)"
+ by (simp add: cong_iff_lin_nat)
-lemma int_cong_add_rcancel: "[(x::int) + a = y + a] (mod n) \<longleftrightarrow> [x = y] (mod n)"
- by (simp add: int_cong_iff_lin)
+lemma cong_add_rcancel_int: "[(x::int) + a = y + a] (mod n) \<longleftrightarrow> [x = y] (mod n)"
+ by (simp add: cong_iff_lin_int)
-lemma nat_cong_add_lcancel_0: "[(a::nat) + x = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"
- by (simp add: nat_cong_iff_lin)
+lemma cong_add_lcancel_0_nat: "[(a::nat) + x = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"
+ by (simp add: cong_iff_lin_nat)
-lemma int_cong_add_lcancel_0: "[(a::int) + x = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"
- by (simp add: int_cong_iff_lin)
+lemma cong_add_lcancel_0_int: "[(a::int) + x = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"
+ by (simp add: cong_iff_lin_int)
-lemma nat_cong_add_rcancel_0: "[x + (a::nat) = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"
- by (simp add: nat_cong_iff_lin)
+lemma cong_add_rcancel_0_nat: "[x + (a::nat) = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"
+ by (simp add: cong_iff_lin_nat)
-lemma int_cong_add_rcancel_0: "[x + (a::int) = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"
- by (simp add: int_cong_iff_lin)
+lemma cong_add_rcancel_0_int: "[x + (a::int) = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"
+ by (simp add: cong_iff_lin_int)
-lemma nat_cong_dvd_modulus: "[(x::nat) = y] (mod m) \<Longrightarrow> n dvd m \<Longrightarrow>
+lemma cong_dvd_modulus_nat: "[(x::nat) = y] (mod m) \<Longrightarrow> n dvd m \<Longrightarrow>
[x = y] (mod n)"
- apply (auto simp add: nat_cong_iff_lin dvd_def)
+ apply (auto simp add: cong_iff_lin_nat dvd_def)
apply (rule_tac x="k1 * k" in exI)
apply (rule_tac x="k2 * k" in exI)
apply (simp add: ring_simps)
done
-lemma int_cong_dvd_modulus: "[(x::int) = y] (mod m) \<Longrightarrow> n dvd m \<Longrightarrow>
+lemma cong_dvd_modulus_int: "[(x::int) = y] (mod m) \<Longrightarrow> n dvd m \<Longrightarrow>
[x = y] (mod n)"
- by (auto simp add: int_cong_altdef dvd_def)
+ by (auto simp add: cong_altdef_int dvd_def)
-lemma nat_cong_dvd_eq: "[(x::nat) = y] (mod n) \<Longrightarrow> n dvd x \<longleftrightarrow> n dvd y"
+lemma cong_dvd_eq_nat: "[(x::nat) = y] (mod n) \<Longrightarrow> n dvd x \<longleftrightarrow> n dvd y"
by (unfold cong_nat_def, auto simp add: dvd_eq_mod_eq_0)
-lemma int_cong_dvd_eq: "[(x::int) = y] (mod n) \<Longrightarrow> n dvd x \<longleftrightarrow> n dvd y"
+lemma cong_dvd_eq_int: "[(x::int) = y] (mod n) \<Longrightarrow> n dvd x \<longleftrightarrow> n dvd y"
by (unfold cong_int_def, auto simp add: dvd_eq_mod_eq_0)
-lemma nat_cong_mod: "(n::nat) ~= 0 \<Longrightarrow> [a mod n = a] (mod n)"
+lemma cong_mod_nat: "(n::nat) ~= 0 \<Longrightarrow> [a mod n = a] (mod n)"
by (simp add: cong_nat_def)
-lemma int_cong_mod: "(n::int) ~= 0 \<Longrightarrow> [a mod n = a] (mod n)"
+lemma cong_mod_int: "(n::int) ~= 0 \<Longrightarrow> [a mod n = a] (mod n)"
by (simp add: cong_int_def)
-lemma nat_mod_mult_cong: "(a::nat) ~= 0 \<Longrightarrow> b ~= 0
+lemma mod_mult_cong_nat: "(a::nat) ~= 0 \<Longrightarrow> b ~= 0
\<Longrightarrow> [x mod (a * b) = y] (mod a) \<longleftrightarrow> [x = y] (mod a)"
by (simp add: cong_nat_def mod_mult2_eq mod_add_left_eq)
-lemma int_neg_cong: "([(a::int) = b] (mod m)) = ([-a = -b] (mod m))"
- apply (simp add: int_cong_altdef)
+lemma neg_cong_int: "([(a::int) = b] (mod m)) = ([-a = -b] (mod m))"
+ apply (simp add: cong_altdef_int)
apply (subst dvd_minus_iff [symmetric])
apply (simp add: ring_simps)
done
-lemma int_cong_modulus_neg: "([(a::int) = b] (mod m)) = ([a = b] (mod -m))"
- by (auto simp add: int_cong_altdef)
+lemma cong_modulus_neg_int: "([(a::int) = b] (mod m)) = ([a = b] (mod -m))"
+ by (auto simp add: cong_altdef_int)
-lemma int_mod_mult_cong: "(a::int) ~= 0 \<Longrightarrow> b ~= 0
+lemma mod_mult_cong_int: "(a::int) ~= 0 \<Longrightarrow> b ~= 0
\<Longrightarrow> [x mod (a * b) = y] (mod a) \<longleftrightarrow> [x = y] (mod a)"
apply (case_tac "b > 0")
apply (simp add: cong_int_def mod_mod_cancel mod_add_left_eq)
- apply (subst (1 2) int_cong_modulus_neg)
+ apply (subst (1 2) cong_modulus_neg_int)
apply (unfold cong_int_def)
apply (subgoal_tac "a * b = (-a * -b)")
apply (erule ssubst)
@@ -577,89 +577,89 @@
apply (auto simp add: mod_add_left_eq)
done
-lemma nat_cong_to_1: "([(a::nat) = 1] (mod n)) \<Longrightarrow> (n dvd (a - 1))"
+lemma cong_to_1_nat: "([(a::nat) = 1] (mod n)) \<Longrightarrow> (n dvd (a - 1))"
apply (case_tac "a = 0")
apply force
- apply (subst (asm) nat_cong_altdef)
+ apply (subst (asm) cong_altdef_nat)
apply auto
done
-lemma nat_0_cong_1: "[(0::nat) = 1] (mod n) = (n = 1)"
+lemma cong_0_1_nat: "[(0::nat) = 1] (mod n) = (n = 1)"
by (unfold cong_nat_def, auto)
-lemma int_0_cong_1: "[(0::int) = 1] (mod n) = ((n = 1) | (n = -1))"
+lemma cong_0_1_int: "[(0::int) = 1] (mod n) = ((n = 1) | (n = -1))"
by (unfold cong_int_def, auto simp add: zmult_eq_1_iff)
-lemma nat_cong_to_1': "[(a::nat) = 1] (mod n) \<longleftrightarrow>
+lemma cong_to_1'_nat: "[(a::nat) = 1] (mod n) \<longleftrightarrow>
a = 0 \<and> n = 1 \<or> (\<exists>m. a = 1 + m * n)"
apply (case_tac "n = 1")
apply auto [1]
apply (drule_tac x = "a - 1" in spec)
apply force
apply (case_tac "a = 0")
- apply (auto simp add: nat_0_cong_1) [1]
+ apply (auto simp add: cong_0_1_nat) [1]
apply (rule iffI)
- apply (drule nat_cong_to_1)
+ apply (drule cong_to_1_nat)
apply (unfold dvd_def)
apply auto [1]
apply (rule_tac x = k in exI)
apply (auto simp add: ring_simps) [1]
- apply (subst nat_cong_altdef)
+ apply (subst cong_altdef_nat)
apply (auto simp add: dvd_def)
done
-lemma nat_cong_le: "(y::nat) <= x \<Longrightarrow> [x = y] (mod n) \<longleftrightarrow> (\<exists>q. x = q * n + y)"
- apply (subst nat_cong_altdef)
+lemma cong_le_nat: "(y::nat) <= x \<Longrightarrow> [x = y] (mod n) \<longleftrightarrow> (\<exists>q. x = q * n + y)"
+ apply (subst cong_altdef_nat)
apply assumption
apply (unfold dvd_def, auto simp add: ring_simps)
apply (rule_tac x = k in exI)
apply auto
done
-lemma nat_cong_solve: "(a::nat) \<noteq> 0 \<Longrightarrow> EX x. [a * x = gcd a n] (mod n)"
+lemma cong_solve_nat: "(a::nat) \<noteq> 0 \<Longrightarrow> EX x. [a * x = gcd a n] (mod n)"
apply (case_tac "n = 0")
apply force
- apply (frule nat_bezout [of a n], auto)
+ apply (frule bezout_nat [of a n], auto)
apply (rule exI, erule ssubst)
- apply (rule nat_cong_trans)
- apply (rule nat_cong_add)
+ apply (rule cong_trans_nat)
+ apply (rule cong_add_nat)
apply (subst mult_commute)
- apply (rule nat_cong_mult_self)
+ apply (rule cong_mult_self_nat)
prefer 2
apply simp
- apply (rule nat_cong_refl)
- apply (rule nat_cong_refl)
+ apply (rule cong_refl_nat)
+ apply (rule cong_refl_nat)
done
-lemma int_cong_solve: "(a::int) \<noteq> 0 \<Longrightarrow> EX x. [a * x = gcd a n] (mod n)"
+lemma cong_solve_int: "(a::int) \<noteq> 0 \<Longrightarrow> EX x. [a * x = gcd a n] (mod n)"
apply (case_tac "n = 0")
apply (case_tac "a \<ge> 0")
apply auto
apply (rule_tac x = "-1" in exI)
apply auto
- apply (insert int_bezout [of a n], auto)
+ apply (insert bezout_int [of a n], auto)
apply (rule exI)
apply (erule subst)
- apply (rule int_cong_trans)
+ apply (rule cong_trans_int)
prefer 2
- apply (rule int_cong_add)
- apply (rule int_cong_refl)
- apply (rule int_cong_sym)
- apply (rule int_cong_mult_self)
+ apply (rule cong_add_int)
+ apply (rule cong_refl_int)
+ apply (rule cong_sym_int)
+ apply (rule cong_mult_self_int)
apply simp
apply (subst mult_commute)
- apply (rule int_cong_refl)
+ apply (rule cong_refl_int)
done
-lemma nat_cong_solve_dvd:
+lemma cong_solve_dvd_nat:
assumes a: "(a::nat) \<noteq> 0" and b: "gcd a n dvd d"
shows "EX x. [a * x = d] (mod n)"
proof -
- from nat_cong_solve [OF a] obtain x where
+ from cong_solve_nat [OF a] obtain x where
"[a * x = gcd a n](mod n)"
by auto
hence "[(d div gcd a n) * (a * x) = (d div gcd a n) * gcd a n] (mod n)"
- by (elim nat_cong_scalar2)
+ by (elim cong_scalar2_nat)
also from b have "(d div gcd a n) * gcd a n = d"
by (rule dvd_div_mult_self)
also have "(d div gcd a n) * (a * x) = a * (d div gcd a n * x)"
@@ -668,15 +668,15 @@
by auto
qed
-lemma int_cong_solve_dvd:
+lemma cong_solve_dvd_int:
assumes a: "(a::int) \<noteq> 0" and b: "gcd a n dvd d"
shows "EX x. [a * x = d] (mod n)"
proof -
- from int_cong_solve [OF a] obtain x where
+ from cong_solve_int [OF a] obtain x where
"[a * x = gcd a n](mod n)"
by auto
hence "[(d div gcd a n) * (a * x) = (d div gcd a n) * gcd a n] (mod n)"
- by (elim int_cong_scalar2)
+ by (elim cong_scalar2_int)
also from b have "(d div gcd a n) * gcd a n = d"
by (rule dvd_div_mult_self)
also have "(d div gcd a n) * (a * x) = a * (d div gcd a n * x)"
@@ -685,41 +685,41 @@
by auto
qed
-lemma nat_cong_solve_coprime: "coprime (a::nat) n \<Longrightarrow>
+lemma cong_solve_coprime_nat: "coprime (a::nat) n \<Longrightarrow>
EX x. [a * x = 1] (mod n)"
apply (case_tac "a = 0")
apply force
- apply (frule nat_cong_solve [of a n])
+ apply (frule cong_solve_nat [of a n])
apply auto
done
-lemma int_cong_solve_coprime: "coprime (a::int) n \<Longrightarrow>
+lemma cong_solve_coprime_int: "coprime (a::int) n \<Longrightarrow>
EX x. [a * x = 1] (mod n)"
apply (case_tac "a = 0")
apply auto
apply (case_tac "n \<ge> 0")
apply auto
apply (subst cong_int_def, auto)
- apply (frule int_cong_solve [of a n])
+ apply (frule cong_solve_int [of a n])
apply auto
done
-lemma nat_coprime_iff_invertible: "m > (1::nat) \<Longrightarrow> coprime a m =
+lemma coprime_iff_invertible_nat: "m > (1::nat) \<Longrightarrow> coprime a m =
(EX x. [a * x = 1] (mod m))"
- apply (auto intro: nat_cong_solve_coprime)
- apply (unfold cong_nat_def, auto intro: nat_invertible_coprime)
+ apply (auto intro: cong_solve_coprime_nat)
+ apply (unfold cong_nat_def, auto intro: invertible_coprime_nat)
done
-lemma int_coprime_iff_invertible: "m > (1::int) \<Longrightarrow> coprime a m =
+lemma coprime_iff_invertible_int: "m > (1::int) \<Longrightarrow> coprime a m =
(EX x. [a * x = 1] (mod m))"
- apply (auto intro: int_cong_solve_coprime)
+ apply (auto intro: cong_solve_coprime_int)
apply (unfold cong_int_def)
- apply (auto intro: int_invertible_coprime)
+ apply (auto intro: invertible_coprime_int)
done
-lemma int_coprime_iff_invertible': "m > (1::int) \<Longrightarrow> coprime a m =
+lemma coprime_iff_invertible'_int: "m > (1::int) \<Longrightarrow> coprime a m =
(EX x. 0 <= x & x < m & [a * x = 1] (mod m))"
- apply (subst int_coprime_iff_invertible)
+ apply (subst coprime_iff_invertible_int)
apply auto
apply (auto simp add: cong_int_def)
apply (rule_tac x = "x mod m" in exI)
@@ -727,200 +727,200 @@
done
-lemma nat_cong_cong_lcm: "[(x::nat) = y] (mod a) \<Longrightarrow>
+lemma cong_cong_lcm_nat: "[(x::nat) = y] (mod a) \<Longrightarrow>
[x = y] (mod b) \<Longrightarrow> [x = y] (mod lcm a b)"
apply (case_tac "y \<le> x")
- apply (auto simp add: nat_cong_altdef nat_lcm_least) [1]
- apply (rule nat_cong_sym)
- apply (subst (asm) (1 2) nat_cong_sym_eq)
- apply (auto simp add: nat_cong_altdef nat_lcm_least)
+ apply (auto simp add: cong_altdef_nat lcm_least_nat) [1]
+ apply (rule cong_sym_nat)
+ apply (subst (asm) (1 2) cong_sym_eq_nat)
+ apply (auto simp add: cong_altdef_nat lcm_least_nat)
done
-lemma int_cong_cong_lcm: "[(x::int) = y] (mod a) \<Longrightarrow>
+lemma cong_cong_lcm_int: "[(x::int) = y] (mod a) \<Longrightarrow>
[x = y] (mod b) \<Longrightarrow> [x = y] (mod lcm a b)"
- by (auto simp add: int_cong_altdef int_lcm_least) [1]
+ by (auto simp add: cong_altdef_int lcm_least_int) [1]
-lemma nat_cong_cong_coprime: "coprime a b \<Longrightarrow> [(x::nat) = y] (mod a) \<Longrightarrow>
+lemma cong_cong_coprime_nat: "coprime a b \<Longrightarrow> [(x::nat) = y] (mod a) \<Longrightarrow>
[x = y] (mod b) \<Longrightarrow> [x = y] (mod a * b)"
- apply (frule (1) nat_cong_cong_lcm)back
+ apply (frule (1) cong_cong_lcm_nat)back
apply (simp add: lcm_nat_def)
done
-lemma int_cong_cong_coprime: "coprime a b \<Longrightarrow> [(x::int) = y] (mod a) \<Longrightarrow>
+lemma cong_cong_coprime_int: "coprime a b \<Longrightarrow> [(x::int) = y] (mod a) \<Longrightarrow>
[x = y] (mod b) \<Longrightarrow> [x = y] (mod a * b)"
- apply (frule (1) int_cong_cong_lcm)back
- apply (simp add: int_lcm_altdef int_cong_abs abs_mult [symmetric])
+ apply (frule (1) cong_cong_lcm_int)back
+ apply (simp add: lcm_altdef_int cong_abs_int abs_mult [symmetric])
done
-lemma nat_cong_cong_setprod_coprime [rule_format]: "finite A \<Longrightarrow>
+lemma cong_cong_setprod_coprime_nat [rule_format]: "finite A \<Longrightarrow>
(ALL i:A. (ALL j:A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))) \<longrightarrow>
(ALL i:A. [(x::nat) = y] (mod m i)) \<longrightarrow>
[x = y] (mod (PROD i:A. m i))"
apply (induct set: finite)
apply auto
- apply (rule nat_cong_cong_coprime)
- apply (subst nat_gcd_commute)
- apply (rule nat_setprod_coprime)
+ apply (rule cong_cong_coprime_nat)
+ apply (subst gcd_commute_nat)
+ apply (rule setprod_coprime_nat)
apply auto
done
-lemma int_cong_cong_setprod_coprime [rule_format]: "finite A \<Longrightarrow>
+lemma cong_cong_setprod_coprime_int [rule_format]: "finite A \<Longrightarrow>
(ALL i:A. (ALL j:A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))) \<longrightarrow>
(ALL i:A. [(x::int) = y] (mod m i)) \<longrightarrow>
[x = y] (mod (PROD i:A. m i))"
apply (induct set: finite)
apply auto
- apply (rule int_cong_cong_coprime)
- apply (subst int_gcd_commute)
- apply (rule int_setprod_coprime)
+ apply (rule cong_cong_coprime_int)
+ apply (subst gcd_commute_int)
+ apply (rule setprod_coprime_int)
apply auto
done
-lemma nat_binary_chinese_remainder_aux:
+lemma binary_chinese_remainder_aux_nat:
assumes a: "coprime (m1::nat) m2"
shows "EX b1 b2. [b1 = 1] (mod m1) \<and> [b1 = 0] (mod m2) \<and>
[b2 = 0] (mod m1) \<and> [b2 = 1] (mod m2)"
proof -
- from nat_cong_solve_coprime [OF a]
+ from cong_solve_coprime_nat [OF a]
obtain x1 where one: "[m1 * x1 = 1] (mod m2)"
by auto
from a have b: "coprime m2 m1"
- by (subst nat_gcd_commute)
- from nat_cong_solve_coprime [OF b]
+ by (subst gcd_commute_nat)
+ from cong_solve_coprime_nat [OF b]
obtain x2 where two: "[m2 * x2 = 1] (mod m1)"
by auto
have "[m1 * x1 = 0] (mod m1)"
- by (subst mult_commute, rule nat_cong_mult_self)
+ by (subst mult_commute, rule cong_mult_self_nat)
moreover have "[m2 * x2 = 0] (mod m2)"
- by (subst mult_commute, rule nat_cong_mult_self)
+ by (subst mult_commute, rule cong_mult_self_nat)
moreover note one two
ultimately show ?thesis by blast
qed
-lemma int_binary_chinese_remainder_aux:
+lemma binary_chinese_remainder_aux_int:
assumes a: "coprime (m1::int) m2"
shows "EX b1 b2. [b1 = 1] (mod m1) \<and> [b1 = 0] (mod m2) \<and>
[b2 = 0] (mod m1) \<and> [b2 = 1] (mod m2)"
proof -
- from int_cong_solve_coprime [OF a]
+ from cong_solve_coprime_int [OF a]
obtain x1 where one: "[m1 * x1 = 1] (mod m2)"
by auto
from a have b: "coprime m2 m1"
- by (subst int_gcd_commute)
- from int_cong_solve_coprime [OF b]
+ by (subst gcd_commute_int)
+ from cong_solve_coprime_int [OF b]
obtain x2 where two: "[m2 * x2 = 1] (mod m1)"
by auto
have "[m1 * x1 = 0] (mod m1)"
- by (subst mult_commute, rule int_cong_mult_self)
+ by (subst mult_commute, rule cong_mult_self_int)
moreover have "[m2 * x2 = 0] (mod m2)"
- by (subst mult_commute, rule int_cong_mult_self)
+ by (subst mult_commute, rule cong_mult_self_int)
moreover note one two
ultimately show ?thesis by blast
qed
-lemma nat_binary_chinese_remainder:
+lemma binary_chinese_remainder_nat:
assumes a: "coprime (m1::nat) m2"
shows "EX x. [x = u1] (mod m1) \<and> [x = u2] (mod m2)"
proof -
- from nat_binary_chinese_remainder_aux [OF a] obtain b1 b2
+ from binary_chinese_remainder_aux_nat [OF a] obtain b1 b2
where "[b1 = 1] (mod m1)" and "[b1 = 0] (mod m2)" and
"[b2 = 0] (mod m1)" and "[b2 = 1] (mod m2)"
by blast
let ?x = "u1 * b1 + u2 * b2"
have "[?x = u1 * 1 + u2 * 0] (mod m1)"
- apply (rule nat_cong_add)
- apply (rule nat_cong_scalar2)
+ apply (rule cong_add_nat)
+ apply (rule cong_scalar2_nat)
apply (rule `[b1 = 1] (mod m1)`)
- apply (rule nat_cong_scalar2)
+ apply (rule cong_scalar2_nat)
apply (rule `[b2 = 0] (mod m1)`)
done
hence "[?x = u1] (mod m1)" by simp
have "[?x = u1 * 0 + u2 * 1] (mod m2)"
- apply (rule nat_cong_add)
- apply (rule nat_cong_scalar2)
+ apply (rule cong_add_nat)
+ apply (rule cong_scalar2_nat)
apply (rule `[b1 = 0] (mod m2)`)
- apply (rule nat_cong_scalar2)
+ apply (rule cong_scalar2_nat)
apply (rule `[b2 = 1] (mod m2)`)
done
hence "[?x = u2] (mod m2)" by simp
with `[?x = u1] (mod m1)` show ?thesis by blast
qed
-lemma int_binary_chinese_remainder:
+lemma binary_chinese_remainder_int:
assumes a: "coprime (m1::int) m2"
shows "EX x. [x = u1] (mod m1) \<and> [x = u2] (mod m2)"
proof -
- from int_binary_chinese_remainder_aux [OF a] obtain b1 b2
+ from binary_chinese_remainder_aux_int [OF a] obtain b1 b2
where "[b1 = 1] (mod m1)" and "[b1 = 0] (mod m2)" and
"[b2 = 0] (mod m1)" and "[b2 = 1] (mod m2)"
by blast
let ?x = "u1 * b1 + u2 * b2"
have "[?x = u1 * 1 + u2 * 0] (mod m1)"
- apply (rule int_cong_add)
- apply (rule int_cong_scalar2)
+ apply (rule cong_add_int)
+ apply (rule cong_scalar2_int)
apply (rule `[b1 = 1] (mod m1)`)
- apply (rule int_cong_scalar2)
+ apply (rule cong_scalar2_int)
apply (rule `[b2 = 0] (mod m1)`)
done
hence "[?x = u1] (mod m1)" by simp
have "[?x = u1 * 0 + u2 * 1] (mod m2)"
- apply (rule int_cong_add)
- apply (rule int_cong_scalar2)
+ apply (rule cong_add_int)
+ apply (rule cong_scalar2_int)
apply (rule `[b1 = 0] (mod m2)`)
- apply (rule int_cong_scalar2)
+ apply (rule cong_scalar2_int)
apply (rule `[b2 = 1] (mod m2)`)
done
hence "[?x = u2] (mod m2)" by simp
with `[?x = u1] (mod m1)` show ?thesis by blast
qed
-lemma nat_cong_modulus_mult: "[(x::nat) = y] (mod m * n) \<Longrightarrow>
+lemma cong_modulus_mult_nat: "[(x::nat) = y] (mod m * n) \<Longrightarrow>
[x = y] (mod m)"
apply (case_tac "y \<le> x")
- apply (simp add: nat_cong_altdef)
+ apply (simp add: cong_altdef_nat)
apply (erule dvd_mult_left)
- apply (rule nat_cong_sym)
- apply (subst (asm) nat_cong_sym_eq)
- apply (simp add: nat_cong_altdef)
+ apply (rule cong_sym_nat)
+ apply (subst (asm) cong_sym_eq_nat)
+ apply (simp add: cong_altdef_nat)
apply (erule dvd_mult_left)
done
-lemma int_cong_modulus_mult: "[(x::int) = y] (mod m * n) \<Longrightarrow>
+lemma cong_modulus_mult_int: "[(x::int) = y] (mod m * n) \<Longrightarrow>
[x = y] (mod m)"
- apply (simp add: int_cong_altdef)
+ apply (simp add: cong_altdef_int)
apply (erule dvd_mult_left)
done
-lemma nat_cong_less_modulus_unique:
+lemma cong_less_modulus_unique_nat:
"[(x::nat) = y] (mod m) \<Longrightarrow> x < m \<Longrightarrow> y < m \<Longrightarrow> x = y"
by (simp add: cong_nat_def)
-lemma nat_binary_chinese_remainder_unique:
+lemma binary_chinese_remainder_unique_nat:
assumes a: "coprime (m1::nat) m2" and
nz: "m1 \<noteq> 0" "m2 \<noteq> 0"
shows "EX! x. x < m1 * m2 \<and> [x = u1] (mod m1) \<and> [x = u2] (mod m2)"
proof -
- from nat_binary_chinese_remainder [OF a] obtain y where
+ from binary_chinese_remainder_nat [OF a] obtain y where
"[y = u1] (mod m1)" and "[y = u2] (mod m2)"
by blast
let ?x = "y mod (m1 * m2)"
from nz have less: "?x < m1 * m2"
by auto
have one: "[?x = u1] (mod m1)"
- apply (rule nat_cong_trans)
+ apply (rule cong_trans_nat)
prefer 2
apply (rule `[y = u1] (mod m1)`)
- apply (rule nat_cong_modulus_mult)
- apply (rule nat_cong_mod)
+ apply (rule cong_modulus_mult_nat)
+ apply (rule cong_mod_nat)
using nz apply auto
done
have two: "[?x = u2] (mod m2)"
- apply (rule nat_cong_trans)
+ apply (rule cong_trans_nat)
prefer 2
apply (rule `[y = u2] (mod m2)`)
apply (subst mult_commute)
- apply (rule nat_cong_modulus_mult)
- apply (rule nat_cong_mod)
+ apply (rule cong_modulus_mult_nat)
+ apply (rule cong_mod_nat)
using nz apply auto
done
have "ALL z. z < m1 * m2 \<and> [z = u1] (mod m1) \<and> [z = u2] (mod m2) \<longrightarrow>
@@ -930,29 +930,29 @@
assume "z < m1 * m2"
assume "[z = u1] (mod m1)" and "[z = u2] (mod m2)"
have "[?x = z] (mod m1)"
- apply (rule nat_cong_trans)
+ apply (rule cong_trans_nat)
apply (rule `[?x = u1] (mod m1)`)
- apply (rule nat_cong_sym)
+ apply (rule cong_sym_nat)
apply (rule `[z = u1] (mod m1)`)
done
moreover have "[?x = z] (mod m2)"
- apply (rule nat_cong_trans)
+ apply (rule cong_trans_nat)
apply (rule `[?x = u2] (mod m2)`)
- apply (rule nat_cong_sym)
+ apply (rule cong_sym_nat)
apply (rule `[z = u2] (mod m2)`)
done
ultimately have "[?x = z] (mod m1 * m2)"
- by (auto intro: nat_coprime_cong_mult a)
+ by (auto intro: coprime_cong_mult_nat a)
with `z < m1 * m2` `?x < m1 * m2` show "z = ?x"
- apply (intro nat_cong_less_modulus_unique)
- apply (auto, erule nat_cong_sym)
+ apply (intro cong_less_modulus_unique_nat)
+ apply (auto, erule cong_sym_nat)
done
qed
with less one two show ?thesis
by auto
qed
-lemma nat_chinese_remainder_aux:
+lemma chinese_remainder_aux_nat:
fixes A :: "'a set" and
m :: "'a \<Rightarrow> nat"
assumes fin: "finite A" and
@@ -963,20 +963,20 @@
fix i
assume "i : A"
with cop have "coprime (PROD j : A - {i}. m j) (m i)"
- by (intro nat_setprod_coprime, auto)
+ by (intro setprod_coprime_nat, auto)
hence "EX x. [(PROD j : A - {i}. m j) * x = 1] (mod m i)"
- by (elim nat_cong_solve_coprime)
+ by (elim cong_solve_coprime_nat)
then obtain x where "[(PROD j : A - {i}. m j) * x = 1] (mod m i)"
by auto
moreover have "[(PROD j : A - {i}. m j) * x = 0]
(mod (PROD j : A - {i}. m j))"
- by (subst mult_commute, rule nat_cong_mult_self)
+ by (subst mult_commute, rule cong_mult_self_nat)
ultimately show "\<exists>a. [a = 1] (mod m i) \<and> [a = 0]
(mod setprod m (A - {i}))"
by blast
qed
-lemma nat_chinese_remainder:
+lemma chinese_remainder_nat:
fixes A :: "'a set" and
m :: "'a \<Rightarrow> nat" and
u :: "'a \<Rightarrow> nat"
@@ -985,7 +985,7 @@
cop: "ALL i:A. (ALL j : A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))"
shows "EX x. (ALL i:A. [x = u i] (mod m i))"
proof -
- from nat_chinese_remainder_aux [OF fin cop] obtain b where
+ from chinese_remainder_aux_nat [OF fin cop] obtain b where
bprop: "ALL i:A. [b i = 1] (mod m i) \<and>
[b i = 0] (mod (PROD j : A - {i}. m j))"
by blast
@@ -1003,13 +1003,13 @@
by auto
also have "[u i * b i + (SUM j:A-{i}. u j * b j) =
u i * 1 + (SUM j:A-{i}. u j * 0)] (mod m i)"
- apply (rule nat_cong_add)
- apply (rule nat_cong_scalar2)
+ apply (rule cong_add_nat)
+ apply (rule cong_scalar2_nat)
using bprop a apply blast
- apply (rule nat_cong_setsum)
- apply (rule nat_cong_scalar2)
+ apply (rule cong_setsum_nat)
+ apply (rule cong_scalar2_nat)
using bprop apply auto
- apply (rule nat_cong_dvd_modulus)
+ apply (rule cong_dvd_modulus_nat)
apply (drule (1) bspec)
apply (erule conjE)
apply assumption
@@ -1022,19 +1022,19 @@
qed
qed
-lemma nat_coprime_cong_prod [rule_format]: "finite A \<Longrightarrow>
+lemma coprime_cong_prod_nat [rule_format]: "finite A \<Longrightarrow>
(ALL i: A. (ALL j: A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))) \<longrightarrow>
(ALL i: A. [(x::nat) = y] (mod m i)) \<longrightarrow>
[x = y] (mod (PROD i:A. m i))"
apply (induct set: finite)
apply auto
- apply (erule (1) nat_coprime_cong_mult)
- apply (subst nat_gcd_commute)
- apply (rule nat_setprod_coprime)
+ apply (erule (1) coprime_cong_mult_nat)
+ apply (subst gcd_commute_nat)
+ apply (rule setprod_coprime_nat)
apply auto
done
-lemma nat_chinese_remainder_unique:
+lemma chinese_remainder_unique_nat:
fixes A :: "'a set" and
m :: "'a \<Rightarrow> nat" and
u :: "'a \<Rightarrow> nat"
@@ -1044,7 +1044,7 @@
cop: "ALL i:A. (ALL j : A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))"
shows "EX! x. x < (PROD i:A. m i) \<and> (ALL i:A. [x = u i] (mod m i))"
proof -
- from nat_chinese_remainder [OF fin cop] obtain y where
+ from chinese_remainder_nat [OF fin cop] obtain y where
one: "(ALL i:A. [y = u i] (mod m i))"
by blast
let ?x = "y mod (PROD i:A. m i)"
@@ -1054,11 +1054,11 @@
by auto
have cong: "ALL i:A. [?x = u i] (mod m i)"
apply auto
- apply (rule nat_cong_trans)
+ apply (rule cong_trans_nat)
prefer 2
using one apply auto
- apply (rule nat_cong_dvd_modulus)
- apply (rule nat_cong_mod)
+ apply (rule cong_dvd_modulus_nat)
+ apply (rule cong_mod_nat)
using prodnz apply auto
apply (rule dvd_setprod)
apply (rule fin)
@@ -1072,16 +1072,16 @@
assume zcong: "(ALL i:A. [z = u i] (mod m i))"
have "ALL i:A. [?x = z] (mod m i)"
apply clarify
- apply (rule nat_cong_trans)
+ apply (rule cong_trans_nat)
using cong apply (erule bspec)
- apply (rule nat_cong_sym)
+ apply (rule cong_sym_nat)
using zcong apply auto
done
with fin cop have "[?x = z] (mod (PROD i:A. m i))"
- by (intro nat_coprime_cong_prod, auto)
+ by (intro coprime_cong_prod_nat, auto)
with zless less show "z = ?x"
- apply (intro nat_cong_less_modulus_unique)
- apply (auto, erule nat_cong_sym)
+ apply (intro cong_less_modulus_unique_nat)
+ apply (auto, erule cong_sym_nat)
done
qed
from less cong unique show ?thesis
--- a/src/HOL/NewNumberTheory/Fib.thy Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/NewNumberTheory/Fib.thy Thu Jul 09 08:55:42 2009 +0200
@@ -86,45 +86,45 @@
subsection {* Fibonacci numbers *}
-lemma nat_fib_0 [simp]: "fib (0::nat) = 0"
+lemma fib_0_nat [simp]: "fib (0::nat) = 0"
by simp
-lemma int_fib_0 [simp]: "fib (0::int) = 0"
+lemma fib_0_int [simp]: "fib (0::int) = 0"
unfolding fib_int_def by simp
-lemma nat_fib_1 [simp]: "fib (1::nat) = 1"
+lemma fib_1_nat [simp]: "fib (1::nat) = 1"
by simp
-lemma nat_fib_Suc_0 [simp]: "fib (Suc 0) = Suc 0"
+lemma fib_Suc_0_nat [simp]: "fib (Suc 0) = Suc 0"
by simp
-lemma int_fib_1 [simp]: "fib (1::int) = 1"
+lemma fib_1_int [simp]: "fib (1::int) = 1"
unfolding fib_int_def by simp
-lemma nat_fib_reduce: "(n::nat) >= 2 \<Longrightarrow> fib n = fib (n - 1) + fib (n - 2)"
+lemma fib_reduce_nat: "(n::nat) >= 2 \<Longrightarrow> fib n = fib (n - 1) + fib (n - 2)"
by simp
declare fib_nat.simps [simp del]
-lemma int_fib_reduce: "(n::int) >= 2 \<Longrightarrow> fib n = fib (n - 1) + fib (n - 2)"
+lemma fib_reduce_int: "(n::int) >= 2 \<Longrightarrow> fib n = fib (n - 1) + fib (n - 2)"
unfolding fib_int_def
- by (auto simp add: nat_fib_reduce nat_diff_distrib)
+ by (auto simp add: fib_reduce_nat nat_diff_distrib)
-lemma int_fib_neg [simp]: "(n::int) < 0 \<Longrightarrow> fib n = 0"
+lemma fib_neg_int [simp]: "(n::int) < 0 \<Longrightarrow> fib n = 0"
unfolding fib_int_def by auto
-lemma nat_fib_2 [simp]: "fib (2::nat) = 1"
- by (subst nat_fib_reduce, auto)
+lemma fib_2_nat [simp]: "fib (2::nat) = 1"
+ by (subst fib_reduce_nat, auto)
-lemma int_fib_2 [simp]: "fib (2::int) = 1"
- by (subst int_fib_reduce, auto)
+lemma fib_2_int [simp]: "fib (2::int) = 1"
+ by (subst fib_reduce_int, auto)
-lemma nat_fib_plus_2: "fib ((n::nat) + 2) = fib (n + 1) + fib n"
- by (subst nat_fib_reduce, auto simp add: One_nat_def)
+lemma fib_plus_2_nat: "fib ((n::nat) + 2) = fib (n + 1) + fib n"
+ by (subst fib_reduce_nat, auto simp add: One_nat_def)
(* the need for One_nat_def is due to the natdiff_cancel_numerals
procedure *)
-lemma nat_fib_induct: "P (0::nat) \<Longrightarrow> P (1::nat) \<Longrightarrow>
+lemma fib_induct_nat: "P (0::nat) \<Longrightarrow> P (1::nat) \<Longrightarrow>
(!!n. P n \<Longrightarrow> P (n + 1) \<Longrightarrow> P (n + 2)) \<Longrightarrow> P n"
apply (atomize, induct n rule: nat_less_induct)
apply auto
@@ -138,13 +138,13 @@
apply (auto simp add: One_nat_def) (* again, natdiff_cancel *)
done
-lemma nat_fib_add: "fib ((n::nat) + k + 1) = fib (k + 1) * fib (n + 1) +
+lemma fib_add_nat: "fib ((n::nat) + k + 1) = fib (k + 1) * fib (n + 1) +
fib k * fib n"
- apply (induct n rule: nat_fib_induct)
+ apply (induct n rule: fib_induct_nat)
apply auto
- apply (subst nat_fib_reduce)
+ apply (subst fib_reduce_nat)
apply (auto simp add: ring_simps)
- apply (subst (1 3 5) nat_fib_reduce)
+ apply (subst (1 3 5) fib_reduce_nat)
apply (auto simp add: ring_simps Suc_eq_plus1)
(* hmmm. Why doesn't "n + (1 + (1 + k))" simplify to "n + k + 2"? *)
apply (subgoal_tac "n + (k + 2) = n + (1 + (1 + k))")
@@ -153,112 +153,112 @@
apply auto
done
-lemma nat_fib_add': "fib (n + Suc k) = fib (Suc k) * fib (Suc n) +
+lemma fib_add'_nat: "fib (n + Suc k) = fib (Suc k) * fib (Suc n) +
fib k * fib n"
- using nat_fib_add by (auto simp add: One_nat_def)
+ using fib_add_nat by (auto simp add: One_nat_def)
(* transfer from nats to ints *)
-lemma int_fib_add [rule_format]: "(n::int) >= 0 \<Longrightarrow> k >= 0 \<Longrightarrow>
+lemma fib_add_int [rule_format]: "(n::int) >= 0 \<Longrightarrow> k >= 0 \<Longrightarrow>
fib (n + k + 1) = fib (k + 1) * fib (n + 1) +
fib k * fib n "
- by (rule nat_fib_add [transferred])
+ by (rule fib_add_nat [transferred])
-lemma nat_fib_neq_0: "(n::nat) > 0 \<Longrightarrow> fib n ~= 0"
- apply (induct n rule: nat_fib_induct)
- apply (auto simp add: nat_fib_plus_2)
+lemma fib_neq_0_nat: "(n::nat) > 0 \<Longrightarrow> fib n ~= 0"
+ apply (induct n rule: fib_induct_nat)
+ apply (auto simp add: fib_plus_2_nat)
done
-lemma nat_fib_gr_0: "(n::nat) > 0 \<Longrightarrow> fib n > 0"
- by (frule nat_fib_neq_0, simp)
+lemma fib_gr_0_nat: "(n::nat) > 0 \<Longrightarrow> fib n > 0"
+ by (frule fib_neq_0_nat, simp)
-lemma int_fib_gr_0: "(n::int) > 0 \<Longrightarrow> fib n > 0"
- unfolding fib_int_def by (simp add: nat_fib_gr_0)
+lemma fib_gr_0_int: "(n::int) > 0 \<Longrightarrow> fib n > 0"
+ unfolding fib_int_def by (simp add: fib_gr_0_nat)
text {*
\medskip Concrete Mathematics, page 278: Cassini's identity. The proof is
much easier using integers, not natural numbers!
*}
-lemma int_fib_Cassini_aux: "fib (int n + 2) * fib (int n) -
+lemma fib_Cassini_aux_int: "fib (int n + 2) * fib (int n) -
(fib (int n + 1))^2 = (-1)^(n + 1)"
apply (induct n)
- apply (auto simp add: ring_simps power2_eq_square int_fib_reduce
+ apply (auto simp add: ring_simps power2_eq_square fib_reduce_int
power_add)
done
-lemma int_fib_Cassini: "n >= 0 \<Longrightarrow> fib (n + 2) * fib n -
+lemma fib_Cassini_int: "n >= 0 \<Longrightarrow> fib (n + 2) * fib n -
(fib (n + 1))^2 = (-1)^(nat n + 1)"
- by (insert int_fib_Cassini_aux [of "nat n"], auto)
+ by (insert fib_Cassini_aux_int [of "nat n"], auto)
(*
-lemma int_fib_Cassini': "n >= 0 \<Longrightarrow> fib (n + 2) * fib n =
+lemma fib_Cassini'_int: "n >= 0 \<Longrightarrow> fib (n + 2) * fib n =
(fib (n + 1))^2 + (-1)^(nat n + 1)"
- by (frule int_fib_Cassini, simp)
+ by (frule fib_Cassini_int, simp)
*)
-lemma int_fib_Cassini': "n >= 0 \<Longrightarrow> fib ((n::int) + 2) * fib n =
+lemma fib_Cassini'_int: "n >= 0 \<Longrightarrow> fib ((n::int) + 2) * fib n =
(if even n then tsub ((fib (n + 1))^2) 1
else (fib (n + 1))^2 + 1)"
- apply (frule int_fib_Cassini, auto simp add: pos_int_even_equiv_nat_even)
+ apply (frule fib_Cassini_int, auto simp add: pos_int_even_equiv_nat_even)
apply (subst tsub_eq)
- apply (insert int_fib_gr_0 [of "n + 1"], force)
+ apply (insert fib_gr_0_int [of "n + 1"], force)
apply auto
done
-lemma nat_fib_Cassini: "fib ((n::nat) + 2) * fib n =
+lemma fib_Cassini_nat: "fib ((n::nat) + 2) * fib n =
(if even n then (fib (n + 1))^2 - 1
else (fib (n + 1))^2 + 1)"
- by (rule int_fib_Cassini' [transferred, of n], auto)
+ by (rule fib_Cassini'_int [transferred, of n], auto)
text {* \medskip Toward Law 6.111 of Concrete Mathematics *}
-lemma nat_coprime_fib_plus_1: "coprime (fib (n::nat)) (fib (n + 1))"
- apply (induct n rule: nat_fib_induct)
+lemma coprime_fib_plus_1_nat: "coprime (fib (n::nat)) (fib (n + 1))"
+ apply (induct n rule: fib_induct_nat)
apply auto
- apply (subst (2) nat_fib_reduce)
+ apply (subst (2) fib_reduce_nat)
apply (auto simp add: Suc_eq_plus1) (* again, natdiff_cancel *)
apply (subst add_commute, auto)
- apply (subst nat_gcd_commute, auto simp add: ring_simps)
+ apply (subst gcd_commute_nat, auto simp add: ring_simps)
done
-lemma nat_coprime_fib_Suc: "coprime (fib n) (fib (Suc n))"
- using nat_coprime_fib_plus_1 by (simp add: One_nat_def)
+lemma coprime_fib_Suc_nat: "coprime (fib n) (fib (Suc n))"
+ using coprime_fib_plus_1_nat by (simp add: One_nat_def)
-lemma int_coprime_fib_plus_1:
+lemma coprime_fib_plus_1_int:
"n >= 0 \<Longrightarrow> coprime (fib (n::int)) (fib (n + 1))"
- by (erule nat_coprime_fib_plus_1 [transferred])
+ by (erule coprime_fib_plus_1_nat [transferred])
-lemma nat_gcd_fib_add: "gcd (fib (m::nat)) (fib (n + m)) = gcd (fib m) (fib n)"
- apply (simp add: nat_gcd_commute [of "fib m"])
- apply (rule nat_cases [of _ m])
+lemma gcd_fib_add_nat: "gcd (fib (m::nat)) (fib (n + m)) = gcd (fib m) (fib n)"
+ apply (simp add: gcd_commute_nat [of "fib m"])
+ apply (rule cases_nat [of _ m])
apply simp
apply (subst add_assoc [symmetric])
- apply (simp add: nat_fib_add)
- apply (subst nat_gcd_commute)
+ apply (simp add: fib_add_nat)
+ apply (subst gcd_commute_nat)
apply (subst mult_commute)
- apply (subst nat_gcd_add_mult)
- apply (subst nat_gcd_commute)
- apply (rule nat_gcd_mult_cancel)
- apply (rule nat_coprime_fib_plus_1)
+ apply (subst gcd_add_mult_nat)
+ apply (subst gcd_commute_nat)
+ apply (rule gcd_mult_cancel_nat)
+ apply (rule coprime_fib_plus_1_nat)
done
-lemma int_gcd_fib_add [rule_format]: "m >= 0 \<Longrightarrow> n >= 0 \<Longrightarrow>
+lemma gcd_fib_add_int [rule_format]: "m >= 0 \<Longrightarrow> n >= 0 \<Longrightarrow>
gcd (fib (m::int)) (fib (n + m)) = gcd (fib m) (fib n)"
- by (erule nat_gcd_fib_add [transferred])
+ by (erule gcd_fib_add_nat [transferred])
-lemma nat_gcd_fib_diff: "(m::nat) \<le> n \<Longrightarrow>
+lemma gcd_fib_diff_nat: "(m::nat) \<le> n \<Longrightarrow>
gcd (fib m) (fib (n - m)) = gcd (fib m) (fib n)"
- by (simp add: nat_gcd_fib_add [symmetric, of _ "n-m"])
+ by (simp add: gcd_fib_add_nat [symmetric, of _ "n-m"])
-lemma int_gcd_fib_diff: "0 <= (m::int) \<Longrightarrow> m \<le> n \<Longrightarrow>
+lemma gcd_fib_diff_int: "0 <= (m::int) \<Longrightarrow> m \<le> n \<Longrightarrow>
gcd (fib m) (fib (n - m)) = gcd (fib m) (fib n)"
- by (simp add: int_gcd_fib_add [symmetric, of _ "n-m"])
+ by (simp add: gcd_fib_add_int [symmetric, of _ "n-m"])
-lemma nat_gcd_fib_mod: "0 < (m::nat) \<Longrightarrow>
+lemma gcd_fib_mod_nat: "0 < (m::nat) \<Longrightarrow>
gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)"
proof (induct n rule: less_induct)
case (less n)
@@ -273,7 +273,7 @@
by (simp add: mod_if [of n]) (insert m_n, auto)
also have "\<dots> = gcd (fib m) (fib (n - m))"
by (simp add: less.hyps diff pos_m)
- also have "\<dots> = gcd (fib m) (fib n)" by (simp add: nat_gcd_fib_diff m_n')
+ also have "\<dots> = gcd (fib m) (fib n)" by (simp add: gcd_fib_diff_nat m_n')
finally show "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)" .
next
case False then show "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)"
@@ -281,39 +281,39 @@
qed
qed
-lemma int_gcd_fib_mod:
+lemma gcd_fib_mod_int:
assumes "0 < (m::int)" and "0 <= n"
shows "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)"
- apply (rule nat_gcd_fib_mod [transferred])
+ apply (rule gcd_fib_mod_nat [transferred])
using prems apply auto
done
-lemma nat_fib_gcd: "fib (gcd (m::nat) n) = gcd (fib m) (fib n)"
+lemma fib_gcd_nat: "fib (gcd (m::nat) n) = gcd (fib m) (fib n)"
-- {* Law 6.111 *}
- apply (induct m n rule: nat_gcd_induct)
- apply (simp_all add: nat_gcd_non_0 nat_gcd_commute nat_gcd_fib_mod)
+ apply (induct m n rule: gcd_nat_induct)
+ apply (simp_all add: gcd_non_0_nat gcd_commute_nat gcd_fib_mod_nat)
done
-lemma int_fib_gcd: "m >= 0 \<Longrightarrow> n >= 0 \<Longrightarrow>
+lemma fib_gcd_int: "m >= 0 \<Longrightarrow> n >= 0 \<Longrightarrow>
fib (gcd (m::int) n) = gcd (fib m) (fib n)"
- by (erule nat_fib_gcd [transferred])
+ by (erule fib_gcd_nat [transferred])
-lemma nat_atMost_plus_one: "{..(k::nat) + 1} = insert (k + 1) {..k}"
+lemma atMost_plus_one_nat: "{..(k::nat) + 1} = insert (k + 1) {..k}"
by auto
-theorem nat_fib_mult_eq_setsum:
+theorem fib_mult_eq_setsum_nat:
"fib ((n::nat) + 1) * fib n = (\<Sum>k \<in> {..n}. fib k * fib k)"
apply (induct n)
- apply (auto simp add: nat_atMost_plus_one nat_fib_plus_2 ring_simps)
+ apply (auto simp add: atMost_plus_one_nat fib_plus_2_nat ring_simps)
done
-theorem nat_fib_mult_eq_setsum':
+theorem fib_mult_eq_setsum'_nat:
"fib (Suc n) * fib n = (\<Sum>k \<in> {..n}. fib k * fib k)"
- using nat_fib_mult_eq_setsum by (simp add: One_nat_def)
+ using fib_mult_eq_setsum_nat by (simp add: One_nat_def)
-theorem int_fib_mult_eq_setsum [rule_format]:
+theorem fib_mult_eq_setsum_int [rule_format]:
"n >= 0 \<Longrightarrow> fib ((n::int) + 1) * fib n = (\<Sum>k \<in> {0..n}. fib k * fib k)"
- by (erule nat_fib_mult_eq_setsum [transferred])
+ by (erule fib_mult_eq_setsum_nat [transferred])
end
--- a/src/HOL/NewNumberTheory/MiscAlgebra.thy Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/NewNumberTheory/MiscAlgebra.thy Thu Jul 09 08:55:42 2009 +0200
@@ -12,7 +12,7 @@
(* finiteness stuff *)
-lemma int_bounded_set1 [intro]: "finite {(x::int). a < x & x < b & P x}"
+lemma bounded_set1_int [intro]: "finite {(x::int). a < x & x < b & P x}"
apply (subgoal_tac "{x. a < x & x < b & P x} <= {a<..<b}")
apply (erule finite_subset)
apply auto
--- a/src/HOL/NewNumberTheory/Residues.thy Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/NewNumberTheory/Residues.thy Thu Jul 09 08:55:42 2009 +0200
@@ -105,10 +105,10 @@
apply auto
apply (subgoal_tac "x ~= 0")
apply auto
- apply (rule int_invertible_coprime)
+ apply (rule invertible_coprime_int)
apply (subgoal_tac "x ~= 0")
apply auto
- apply (subst (asm) int_coprime_iff_invertible')
+ apply (subst (asm) coprime_iff_invertible'_int)
apply (rule m_gt_one)
apply (auto simp add: cong_int_def mult_commute)
done
@@ -202,8 +202,8 @@
apply (subgoal_tac "a mod m ~= 0")
apply arith
apply auto
- apply (subst (asm) int_gcd_red)
- apply (subst int_gcd_commute, assumption)
+ apply (subst (asm) gcd_red_int)
+ apply (subst gcd_commute_int, assumption)
done
lemma res_eq_to_cong: "((a mod m) = (b mod m)) = [a = b] (mod (m::int))"
@@ -254,8 +254,8 @@
res_units_eq)
apply (rule classical)
apply (erule notE)
- apply (subst int_gcd_commute)
- apply (rule int_prime_imp_coprime)
+ apply (subst gcd_commute_int)
+ apply (rule prime_imp_coprime_int)
apply (rule p_prime)
apply (rule notI)
apply (frule zdvd_imp_le)
@@ -265,8 +265,8 @@
lemma res_prime_units_eq: "Units R = {1..p - 1}"
apply (subst res_units_eq)
apply auto
- apply (subst int_gcd_commute)
- apply (rule int_prime_imp_coprime)
+ apply (subst gcd_commute_int)
+ apply (rule prime_imp_coprime_int)
apply (rule p_prime)
apply (rule zdvd_not_zless)
apply auto
@@ -367,8 +367,8 @@
apply (intro euler_theorem)
(* auto should get this next part. matching across
substitutions is needed. *)
- apply (frule int_prime_gt_1, arith)
- apply (subst int_gcd_commute, erule int_prime_imp_coprime, assumption)
+ apply (frule prime_gt_1_int, arith)
+ apply (subst gcd_commute_int, erule prime_imp_coprime_int, assumption)
done
also have "phi p = nat p - 1"
by (rule phi_prime, rule prems)
@@ -442,7 +442,7 @@
apply auto
done
also have "\<dots> = fact (p - 1) mod p"
- apply (subst int_fact_altdef)
+ apply (subst fact_altdef_int)
apply (insert prems, force)
apply (subst res_prime_units_eq, rule refl)
done
@@ -452,9 +452,9 @@
qed
lemma wilson_theorem: "prime (p::int) \<Longrightarrow> [fact (p - 1) = - 1] (mod p)"
- apply (frule int_prime_gt_1)
+ apply (frule prime_gt_1_int)
apply (case_tac "p = 2")
- apply (subst int_fact_altdef, simp)
+ apply (subst fact_altdef_int, simp)
apply (subst cong_int_def)
apply simp
apply (rule residues_prime.wilson_theorem1)
--- a/src/HOL/NewNumberTheory/UniqueFactorization.thy Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/NewNumberTheory/UniqueFactorization.thy Thu Jul 09 08:55:42 2009 +0200
@@ -134,7 +134,7 @@
moreover
{
assume "n > 1" and "~ prime n"
- from prems nat_not_prime_eq_prod
+ from prems not_prime_eq_prod_nat
obtain m k where "n = m * k & 1 < m & m < n & 1 < k & k < n"
by blast
with ih obtain Q R where "(ALL p : set_of Q. prime p) & m = (PROD i:#Q. i)"
@@ -182,13 +182,13 @@
a^(count N a) * (PROD i : (set_of N - {a}). i ^ (count N i))".
moreover have "coprime (a ^ count M a)
(PROD i : (set_of N - {a}). i ^ (count N i))"
- apply (subst nat_gcd_commute)
- apply (rule nat_setprod_coprime)
- apply (rule nat_primes_imp_powers_coprime)
+ apply (subst gcd_commute_nat)
+ apply (rule setprod_coprime_nat)
+ apply (rule primes_imp_powers_coprime_nat)
apply (insert prems, auto)
done
ultimately have "a ^ count M a dvd a^(count N a)"
- by (elim nat_coprime_dvd_mult)
+ by (elim coprime_dvd_mult_nat)
with a show ?thesis
by (intro power_dvd_imp_le, auto)
next
@@ -322,66 +322,66 @@
subsection {* Properties of prime factors and multiplicity for nats and ints *}
-lemma int_prime_factors_ge_0 [elim]: "p : prime_factors (n::int) \<Longrightarrow> p >= 0"
+lemma prime_factors_ge_0_int [elim]: "p : prime_factors (n::int) \<Longrightarrow> p >= 0"
by (unfold prime_factors_int_def, auto)
-lemma nat_prime_factors_prime [intro]: "p : prime_factors (n::nat) \<Longrightarrow> prime p"
+lemma prime_factors_prime_nat [intro]: "p : prime_factors (n::nat) \<Longrightarrow> prime p"
apply (case_tac "n = 0")
apply (simp add: prime_factors_nat_def multiset_prime_factorization_def)
apply (auto simp add: prime_factors_nat_def multiset_prime_factorization)
done
-lemma int_prime_factors_prime [intro]:
+lemma prime_factors_prime_int [intro]:
assumes "n >= 0" and "p : prime_factors (n::int)"
shows "prime p"
- apply (rule nat_prime_factors_prime [transferred, of n p])
+ apply (rule prime_factors_prime_nat [transferred, of n p])
using prems apply auto
done
-lemma nat_prime_factors_gt_0 [elim]: "p : prime_factors x \<Longrightarrow> p > (0::nat)"
- by (frule nat_prime_factors_prime, auto)
+lemma prime_factors_gt_0_nat [elim]: "p : prime_factors x \<Longrightarrow> p > (0::nat)"
+ by (frule prime_factors_prime_nat, auto)
-lemma int_prime_factors_gt_0 [elim]: "x >= 0 \<Longrightarrow> p : prime_factors x \<Longrightarrow>
+lemma prime_factors_gt_0_int [elim]: "x >= 0 \<Longrightarrow> p : prime_factors x \<Longrightarrow>
p > (0::int)"
- by (frule (1) int_prime_factors_prime, auto)
+ by (frule (1) prime_factors_prime_int, auto)
-lemma nat_prime_factors_finite [iff]: "finite (prime_factors (n::nat))"
+lemma prime_factors_finite_nat [iff]: "finite (prime_factors (n::nat))"
by (unfold prime_factors_nat_def, auto)
-lemma int_prime_factors_finite [iff]: "finite (prime_factors (n::int))"
+lemma prime_factors_finite_int [iff]: "finite (prime_factors (n::int))"
by (unfold prime_factors_int_def, auto)
-lemma nat_prime_factors_altdef: "prime_factors (n::nat) =
+lemma prime_factors_altdef_nat: "prime_factors (n::nat) =
{p. multiplicity p n > 0}"
by (force simp add: prime_factors_nat_def multiplicity_nat_def)
-lemma int_prime_factors_altdef: "prime_factors (n::int) =
+lemma prime_factors_altdef_int: "prime_factors (n::int) =
{p. p >= 0 & multiplicity p n > 0}"
apply (unfold prime_factors_int_def multiplicity_int_def)
- apply (subst nat_prime_factors_altdef)
+ apply (subst prime_factors_altdef_nat)
apply (auto simp add: image_def)
done
-lemma nat_prime_factorization: "(n::nat) > 0 \<Longrightarrow>
+lemma prime_factorization_nat: "(n::nat) > 0 \<Longrightarrow>
n = (PROD p : prime_factors n. p^(multiplicity p n))"
by (frule multiset_prime_factorization,
simp add: prime_factors_nat_def multiplicity_nat_def msetprod_def)
-thm nat_prime_factorization [transferred]
+thm prime_factorization_nat [transferred]
-lemma int_prime_factorization:
+lemma prime_factorization_int:
assumes "(n::int) > 0"
shows "n = (PROD p : prime_factors n. p^(multiplicity p n))"
- apply (rule nat_prime_factorization [transferred, of n])
+ apply (rule prime_factorization_nat [transferred, of n])
using prems apply auto
done
-lemma nat_neq_zero_eq_gt_zero: "((x::nat) ~= 0) = (x > 0)"
+lemma neq_zero_eq_gt_zero_nat: "((x::nat) ~= 0) = (x > 0)"
by auto
-lemma nat_prime_factorization_unique:
+lemma prime_factorization_unique_nat:
"S = { (p::nat) . f p > 0} \<Longrightarrow> finite S \<Longrightarrow> (ALL p : S. prime p) \<Longrightarrow>
n = (PROD p : S. p^(f p)) \<Longrightarrow>
S = prime_factors n & (ALL p. f p = multiplicity p n)"
@@ -407,23 +407,23 @@
unfolding multiset_def apply force
done
-lemma nat_prime_factors_characterization: "S = {p. 0 < f (p::nat)} \<Longrightarrow>
+lemma prime_factors_characterization_nat: "S = {p. 0 < f (p::nat)} \<Longrightarrow>
finite S \<Longrightarrow> (ALL p:S. prime p) \<Longrightarrow> n = (PROD p:S. p ^ f p) \<Longrightarrow>
prime_factors n = S"
- by (rule nat_prime_factorization_unique [THEN conjunct1, symmetric],
+ by (rule prime_factorization_unique_nat [THEN conjunct1, symmetric],
assumption+)
-lemma nat_prime_factors_characterization':
+lemma prime_factors_characterization'_nat:
"finite {p. 0 < f (p::nat)} \<Longrightarrow>
(ALL p. 0 < f p \<longrightarrow> prime p) \<Longrightarrow>
prime_factors (PROD p | 0 < f p . p ^ f p) = {p. 0 < f p}"
- apply (rule nat_prime_factors_characterization)
+ apply (rule prime_factors_characterization_nat)
apply auto
done
(* A minor glitch:*)
-thm nat_prime_factors_characterization'
+thm prime_factors_characterization'_nat
[where f = "%x. f (int (x::nat))",
transferred direction: nat "op <= (0::int)", rule_format]
@@ -432,106 +432,106 @@
remain a comparison between nats. But the transfer still works.
*)
-lemma int_primes_characterization' [rule_format]:
+lemma primes_characterization'_int [rule_format]:
"finite {p. p >= 0 & 0 < f (p::int)} \<Longrightarrow>
(ALL p. 0 < f p \<longrightarrow> prime p) \<Longrightarrow>
prime_factors (PROD p | p >=0 & 0 < f p . p ^ f p) =
{p. p >= 0 & 0 < f p}"
- apply (insert nat_prime_factors_characterization'
+ apply (insert prime_factors_characterization'_nat
[where f = "%x. f (int (x::nat))",
transferred direction: nat "op <= (0::int)"])
apply auto
done
-lemma int_prime_factors_characterization: "S = {p. 0 < f (p::int)} \<Longrightarrow>
+lemma prime_factors_characterization_int: "S = {p. 0 < f (p::int)} \<Longrightarrow>
finite S \<Longrightarrow> (ALL p:S. prime p) \<Longrightarrow> n = (PROD p:S. p ^ f p) \<Longrightarrow>
prime_factors n = S"
apply simp
apply (subgoal_tac "{p. 0 < f p} = {p. 0 <= p & 0 < f p}")
apply (simp only:)
- apply (subst int_primes_characterization')
+ apply (subst primes_characterization'_int)
apply auto
- apply (auto simp add: int_prime_ge_0)
+ apply (auto simp add: prime_ge_0_int)
done
-lemma nat_multiplicity_characterization: "S = {p. 0 < f (p::nat)} \<Longrightarrow>
+lemma multiplicity_characterization_nat: "S = {p. 0 < f (p::nat)} \<Longrightarrow>
finite S \<Longrightarrow> (ALL p:S. prime p) \<Longrightarrow> n = (PROD p:S. p ^ f p) \<Longrightarrow>
multiplicity p n = f p"
- by (frule nat_prime_factorization_unique [THEN conjunct2, rule_format,
+ by (frule prime_factorization_unique_nat [THEN conjunct2, rule_format,
symmetric], auto)
-lemma nat_multiplicity_characterization': "finite {p. 0 < f (p::nat)} \<longrightarrow>
+lemma multiplicity_characterization'_nat: "finite {p. 0 < f (p::nat)} \<longrightarrow>
(ALL p. 0 < f p \<longrightarrow> prime p) \<longrightarrow>
multiplicity p (PROD p | 0 < f p . p ^ f p) = f p"
apply (rule impI)+
- apply (rule nat_multiplicity_characterization)
+ apply (rule multiplicity_characterization_nat)
apply auto
done
-lemma int_multiplicity_characterization' [rule_format]:
+lemma multiplicity_characterization'_int [rule_format]:
"finite {p. p >= 0 & 0 < f (p::int)} \<Longrightarrow>
(ALL p. 0 < f p \<longrightarrow> prime p) \<Longrightarrow> p >= 0 \<Longrightarrow>
multiplicity p (PROD p | p >= 0 & 0 < f p . p ^ f p) = f p"
- apply (insert nat_multiplicity_characterization'
+ apply (insert multiplicity_characterization'_nat
[where f = "%x. f (int (x::nat))",
transferred direction: nat "op <= (0::int)", rule_format])
apply auto
done
-lemma int_multiplicity_characterization: "S = {p. 0 < f (p::int)} \<Longrightarrow>
+lemma multiplicity_characterization_int: "S = {p. 0 < f (p::int)} \<Longrightarrow>
finite S \<Longrightarrow> (ALL p:S. prime p) \<Longrightarrow> n = (PROD p:S. p ^ f p) \<Longrightarrow>
p >= 0 \<Longrightarrow> multiplicity p n = f p"
apply simp
apply (subgoal_tac "{p. 0 < f p} = {p. 0 <= p & 0 < f p}")
apply (simp only:)
- apply (subst int_multiplicity_characterization')
+ apply (subst multiplicity_characterization'_int)
apply auto
- apply (auto simp add: int_prime_ge_0)
+ apply (auto simp add: prime_ge_0_int)
done
-lemma nat_multiplicity_zero [simp]: "multiplicity (p::nat) 0 = 0"
+lemma multiplicity_zero_nat [simp]: "multiplicity (p::nat) 0 = 0"
by (simp add: multiplicity_nat_def multiset_prime_factorization_def)
-lemma int_multiplicity_zero [simp]: "multiplicity (p::int) 0 = 0"
+lemma multiplicity_zero_int [simp]: "multiplicity (p::int) 0 = 0"
by (simp add: multiplicity_int_def)
-lemma nat_multiplicity_one [simp]: "multiplicity p (1::nat) = 0"
- by (subst nat_multiplicity_characterization [where f = "%x. 0"], auto)
+lemma multiplicity_one_nat [simp]: "multiplicity p (1::nat) = 0"
+ by (subst multiplicity_characterization_nat [where f = "%x. 0"], auto)
-lemma int_multiplicity_one [simp]: "multiplicity p (1::int) = 0"
+lemma multiplicity_one_int [simp]: "multiplicity p (1::int) = 0"
by (simp add: multiplicity_int_def)
-lemma nat_multiplicity_prime [simp]: "prime (p::nat) \<Longrightarrow> multiplicity p p = 1"
- apply (subst nat_multiplicity_characterization
+lemma multiplicity_prime_nat [simp]: "prime (p::nat) \<Longrightarrow> multiplicity p p = 1"
+ apply (subst multiplicity_characterization_nat
[where f = "(%q. if q = p then 1 else 0)"])
apply auto
apply (case_tac "x = p")
apply auto
done
-lemma int_multiplicity_prime [simp]: "prime (p::int) \<Longrightarrow> multiplicity p p = 1"
+lemma multiplicity_prime_int [simp]: "prime (p::int) \<Longrightarrow> multiplicity p p = 1"
unfolding prime_int_def multiplicity_int_def by auto
-lemma nat_multiplicity_prime_power [simp]: "prime (p::nat) \<Longrightarrow>
+lemma multiplicity_prime_power_nat [simp]: "prime (p::nat) \<Longrightarrow>
multiplicity p (p^n) = n"
apply (case_tac "n = 0")
apply auto
- apply (subst nat_multiplicity_characterization
+ apply (subst multiplicity_characterization_nat
[where f = "(%q. if q = p then n else 0)"])
apply auto
apply (case_tac "x = p")
apply auto
done
-lemma int_multiplicity_prime_power [simp]: "prime (p::int) \<Longrightarrow>
+lemma multiplicity_prime_power_int [simp]: "prime (p::int) \<Longrightarrow>
multiplicity p (p^n) = n"
- apply (frule int_prime_ge_0)
+ apply (frule prime_ge_0_int)
apply (auto simp add: prime_int_def multiplicity_int_def nat_power_eq)
done
-lemma nat_multiplicity_nonprime [simp]: "~ prime (p::nat) \<Longrightarrow>
+lemma multiplicity_nonprime_nat [simp]: "~ prime (p::nat) \<Longrightarrow>
multiplicity p n = 0"
apply (case_tac "n = 0")
apply auto
@@ -539,22 +539,22 @@
apply (auto simp add: set_of_def multiplicity_nat_def)
done
-lemma int_multiplicity_nonprime [simp]: "~ prime (p::int) \<Longrightarrow> multiplicity p n = 0"
+lemma multiplicity_nonprime_int [simp]: "~ prime (p::int) \<Longrightarrow> multiplicity p n = 0"
by (unfold multiplicity_int_def prime_int_def, auto)
-lemma nat_multiplicity_not_factor [simp]:
+lemma multiplicity_not_factor_nat [simp]:
"p ~: prime_factors (n::nat) \<Longrightarrow> multiplicity p n = 0"
- by (subst (asm) nat_prime_factors_altdef, auto)
+ by (subst (asm) prime_factors_altdef_nat, auto)
-lemma int_multiplicity_not_factor [simp]:
+lemma multiplicity_not_factor_int [simp]:
"p >= 0 \<Longrightarrow> p ~: prime_factors (n::int) \<Longrightarrow> multiplicity p n = 0"
- by (subst (asm) int_prime_factors_altdef, auto)
+ by (subst (asm) prime_factors_altdef_int, auto)
-lemma nat_multiplicity_product_aux: "(k::nat) > 0 \<Longrightarrow> l > 0 \<Longrightarrow>
+lemma multiplicity_product_aux_nat: "(k::nat) > 0 \<Longrightarrow> l > 0 \<Longrightarrow>
(prime_factors k) Un (prime_factors l) = prime_factors (k * l) &
(ALL p. multiplicity p k + multiplicity p l = multiplicity p (k * l))"
- apply (rule nat_prime_factorization_unique)
- apply (simp only: nat_prime_factors_altdef)
+ apply (rule prime_factorization_unique_nat)
+ apply (simp only: prime_factors_altdef_nat)
apply auto
apply (subst power_add)
apply (subst setprod_timesf)
@@ -568,7 +568,7 @@
(\<Prod>p\<in>prime_factors l - prime_factors k. 1)")
apply (erule ssubst)
apply (simp add: setprod_1)
- apply (erule nat_prime_factorization)
+ apply (erule prime_factorization_nat)
apply (rule setprod_cong, auto)
apply (subgoal_tac "prime_factors k Un prime_factors l = prime_factors l Un
(prime_factors k - prime_factors l)")
@@ -579,47 +579,47 @@
(\<Prod>p\<in>prime_factors k - prime_factors l. 1)")
apply (erule ssubst)
apply (simp add: setprod_1)
- apply (erule nat_prime_factorization)
+ apply (erule prime_factorization_nat)
apply (rule setprod_cong, auto)
done
(* transfer doesn't have the same problem here with the right
choice of rules. *)
-lemma int_multiplicity_product_aux:
+lemma multiplicity_product_aux_int:
assumes "(k::int) > 0" and "l > 0"
shows
"(prime_factors k) Un (prime_factors l) = prime_factors (k * l) &
(ALL p >= 0. multiplicity p k + multiplicity p l = multiplicity p (k * l))"
- apply (rule nat_multiplicity_product_aux [transferred, of l k])
+ apply (rule multiplicity_product_aux_nat [transferred, of l k])
using prems apply auto
done
-lemma nat_prime_factors_product: "(k::nat) > 0 \<Longrightarrow> l > 0 \<Longrightarrow> prime_factors (k * l) =
+lemma prime_factors_product_nat: "(k::nat) > 0 \<Longrightarrow> l > 0 \<Longrightarrow> prime_factors (k * l) =
prime_factors k Un prime_factors l"
- by (rule nat_multiplicity_product_aux [THEN conjunct1, symmetric])
+ by (rule multiplicity_product_aux_nat [THEN conjunct1, symmetric])
-lemma int_prime_factors_product: "(k::int) > 0 \<Longrightarrow> l > 0 \<Longrightarrow> prime_factors (k * l) =
+lemma prime_factors_product_int: "(k::int) > 0 \<Longrightarrow> l > 0 \<Longrightarrow> prime_factors (k * l) =
prime_factors k Un prime_factors l"
- by (rule int_multiplicity_product_aux [THEN conjunct1, symmetric])
+ by (rule multiplicity_product_aux_int [THEN conjunct1, symmetric])
-lemma nat_multiplicity_product: "(k::nat) > 0 \<Longrightarrow> l > 0 \<Longrightarrow> multiplicity p (k * l) =
+lemma multiplicity_product_nat: "(k::nat) > 0 \<Longrightarrow> l > 0 \<Longrightarrow> multiplicity p (k * l) =
multiplicity p k + multiplicity p l"
- by (rule nat_multiplicity_product_aux [THEN conjunct2, rule_format,
+ by (rule multiplicity_product_aux_nat [THEN conjunct2, rule_format,
symmetric])
-lemma int_multiplicity_product: "(k::int) > 0 \<Longrightarrow> l > 0 \<Longrightarrow> p >= 0 \<Longrightarrow>
+lemma multiplicity_product_int: "(k::int) > 0 \<Longrightarrow> l > 0 \<Longrightarrow> p >= 0 \<Longrightarrow>
multiplicity p (k * l) = multiplicity p k + multiplicity p l"
- by (rule int_multiplicity_product_aux [THEN conjunct2, rule_format,
+ by (rule multiplicity_product_aux_int [THEN conjunct2, rule_format,
symmetric])
-lemma nat_multiplicity_setprod: "finite S \<Longrightarrow> (ALL x : S. f x > 0) \<Longrightarrow>
+lemma multiplicity_setprod_nat: "finite S \<Longrightarrow> (ALL x : S. f x > 0) \<Longrightarrow>
multiplicity (p::nat) (PROD x : S. f x) =
(SUM x : S. multiplicity p (f x))"
apply (induct set: finite)
apply auto
- apply (subst nat_multiplicity_product)
+ apply (subst multiplicity_product_nat)
apply auto
done
@@ -643,12 +643,12 @@
add return: transfer_nat_int_sum_prod_closure3
del: transfer_nat_int_sum_prod2 (1)]
-lemma int_multiplicity_setprod: "p >= 0 \<Longrightarrow> finite S \<Longrightarrow>
+lemma multiplicity_setprod_int: "p >= 0 \<Longrightarrow> finite S \<Longrightarrow>
(ALL x : S. f x > 0) \<Longrightarrow>
multiplicity (p::int) (PROD x : S. f x) =
(SUM x : S. multiplicity p (f x))"
- apply (frule nat_multiplicity_setprod
+ apply (frule multiplicity_setprod_nat
[where f = "%x. nat(int(nat(f x)))",
transferred direction: nat "op <= (0::int)"])
apply auto
@@ -663,13 +663,13 @@
declare TransferMorphism_nat_int[transfer
add return: transfer_nat_int_sum_prod2 (1)]
-lemma nat_multiplicity_prod_prime_powers:
+lemma multiplicity_prod_prime_powers_nat:
"finite S \<Longrightarrow> (ALL p : S. prime (p::nat)) \<Longrightarrow>
multiplicity p (PROD p : S. p ^ f p) = (if p : S then f p else 0)"
apply (subgoal_tac "(PROD p : S. p ^ f p) =
(PROD p : S. p ^ (%x. if x : S then f x else 0) p)")
apply (erule ssubst)
- apply (subst nat_multiplicity_characterization)
+ apply (subst multiplicity_characterization_nat)
prefer 5 apply (rule refl)
apply (rule refl)
apply auto
@@ -683,85 +683,85 @@
(* Here the issue with transfer is the implicit quantifier over S *)
-lemma int_multiplicity_prod_prime_powers:
+lemma multiplicity_prod_prime_powers_int:
"(p::int) >= 0 \<Longrightarrow> finite S \<Longrightarrow> (ALL p : S. prime p) \<Longrightarrow>
multiplicity p (PROD p : S. p ^ f p) = (if p : S then f p else 0)"
apply (subgoal_tac "int ` nat ` S = S")
- apply (frule nat_multiplicity_prod_prime_powers [where f = "%x. f(int x)"
+ apply (frule multiplicity_prod_prime_powers_nat [where f = "%x. f(int x)"
and S = "nat ` S", transferred])
apply auto
apply (subst prime_int_def [symmetric])
apply auto
apply (subgoal_tac "xb >= 0")
apply force
- apply (rule int_prime_ge_0)
+ apply (rule prime_ge_0_int)
apply force
apply (subst transfer_nat_int_set_return_embed)
apply (unfold nat_set_def, auto)
done
-lemma nat_multiplicity_distinct_prime_power: "prime (p::nat) \<Longrightarrow> prime q \<Longrightarrow>
+lemma multiplicity_distinct_prime_power_nat: "prime (p::nat) \<Longrightarrow> prime q \<Longrightarrow>
p ~= q \<Longrightarrow> multiplicity p (q^n) = 0"
apply (subgoal_tac "q^n = setprod (%x. x^n) {q}")
apply (erule ssubst)
- apply (subst nat_multiplicity_prod_prime_powers)
+ apply (subst multiplicity_prod_prime_powers_nat)
apply auto
done
-lemma int_multiplicity_distinct_prime_power: "prime (p::int) \<Longrightarrow> prime q \<Longrightarrow>
+lemma multiplicity_distinct_prime_power_int: "prime (p::int) \<Longrightarrow> prime q \<Longrightarrow>
p ~= q \<Longrightarrow> multiplicity p (q^n) = 0"
- apply (frule int_prime_ge_0 [of q])
- apply (frule nat_multiplicity_distinct_prime_power [transferred leaving: n])
+ apply (frule prime_ge_0_int [of q])
+ apply (frule multiplicity_distinct_prime_power_nat [transferred leaving: n])
prefer 4
apply assumption
apply auto
done
-lemma nat_dvd_multiplicity:
+lemma dvd_multiplicity_nat:
"(0::nat) < y \<Longrightarrow> x dvd y \<Longrightarrow> multiplicity p x <= multiplicity p y"
apply (case_tac "x = 0")
- apply (auto simp add: dvd_def nat_multiplicity_product)
+ apply (auto simp add: dvd_def multiplicity_product_nat)
done
-lemma int_dvd_multiplicity:
+lemma dvd_multiplicity_int:
"(0::int) < y \<Longrightarrow> 0 <= x \<Longrightarrow> x dvd y \<Longrightarrow> p >= 0 \<Longrightarrow>
multiplicity p x <= multiplicity p y"
apply (case_tac "x = 0")
apply (auto simp add: dvd_def)
apply (subgoal_tac "0 < k")
- apply (auto simp add: int_multiplicity_product)
+ apply (auto simp add: multiplicity_product_int)
apply (erule zero_less_mult_pos)
apply arith
done
-lemma nat_dvd_prime_factors [intro]:
+lemma dvd_prime_factors_nat [intro]:
"0 < (y::nat) \<Longrightarrow> x dvd y \<Longrightarrow> prime_factors x <= prime_factors y"
- apply (simp only: nat_prime_factors_altdef)
+ apply (simp only: prime_factors_altdef_nat)
apply auto
- apply (frule nat_dvd_multiplicity)
+ apply (frule dvd_multiplicity_nat)
apply auto
(* It is a shame that auto and arith don't get this. *)
apply (erule order_less_le_trans)back
apply assumption
done
-lemma int_dvd_prime_factors [intro]:
+lemma dvd_prime_factors_int [intro]:
"0 < (y::int) \<Longrightarrow> 0 <= x \<Longrightarrow> x dvd y \<Longrightarrow> prime_factors x <= prime_factors y"
- apply (auto simp add: int_prime_factors_altdef)
+ apply (auto simp add: prime_factors_altdef_int)
apply (erule order_less_le_trans)
- apply (rule int_dvd_multiplicity)
+ apply (rule dvd_multiplicity_int)
apply auto
done
-lemma nat_multiplicity_dvd: "0 < (x::nat) \<Longrightarrow> 0 < y \<Longrightarrow>
+lemma multiplicity_dvd_nat: "0 < (x::nat) \<Longrightarrow> 0 < y \<Longrightarrow>
ALL p. multiplicity p x <= multiplicity p y \<Longrightarrow>
x dvd y"
- apply (subst nat_prime_factorization [of x], assumption)
- apply (subst nat_prime_factorization [of y], assumption)
+ apply (subst prime_factorization_nat [of x], assumption)
+ apply (subst prime_factorization_nat [of y], assumption)
apply (rule setprod_dvd_setprod_subset2)
apply force
- apply (subst nat_prime_factors_altdef)+
+ apply (subst prime_factors_altdef_nat)+
apply auto
(* Again, a shame that auto and arith don't get this. *)
apply (drule_tac x = xa in spec, auto)
@@ -769,14 +769,14 @@
apply blast
done
-lemma int_multiplicity_dvd: "0 < (x::int) \<Longrightarrow> 0 < y \<Longrightarrow>
+lemma multiplicity_dvd_int: "0 < (x::int) \<Longrightarrow> 0 < y \<Longrightarrow>
ALL p >= 0. multiplicity p x <= multiplicity p y \<Longrightarrow>
x dvd y"
- apply (subst int_prime_factorization [of x], assumption)
- apply (subst int_prime_factorization [of y], assumption)
+ apply (subst prime_factorization_int [of x], assumption)
+ apply (subst prime_factorization_int [of y], assumption)
apply (rule setprod_dvd_setprod_subset2)
apply force
- apply (subst int_prime_factors_altdef)+
+ apply (subst prime_factors_altdef_int)+
apply auto
apply (rule dvd_power_le)
apply auto
@@ -785,83 +785,83 @@
apply auto
done
-lemma nat_multiplicity_dvd': "(0::nat) < x \<Longrightarrow>
+lemma multiplicity_dvd'_nat: "(0::nat) < x \<Longrightarrow>
\<forall>p. prime p \<longrightarrow> multiplicity p x \<le> multiplicity p y \<Longrightarrow> x dvd y"
apply (cases "y = 0")
apply auto
- apply (rule nat_multiplicity_dvd, auto)
+ apply (rule multiplicity_dvd_nat, auto)
apply (case_tac "prime p")
apply auto
done
-lemma int_multiplicity_dvd': "(0::int) < x \<Longrightarrow> 0 <= y \<Longrightarrow>
+lemma multiplicity_dvd'_int: "(0::int) < x \<Longrightarrow> 0 <= y \<Longrightarrow>
\<forall>p. prime p \<longrightarrow> multiplicity p x \<le> multiplicity p y \<Longrightarrow> x dvd y"
apply (cases "y = 0")
apply auto
- apply (rule int_multiplicity_dvd, auto)
+ apply (rule multiplicity_dvd_int, auto)
apply (case_tac "prime p")
apply auto
done
-lemma nat_dvd_multiplicity_eq: "0 < (x::nat) \<Longrightarrow> 0 < y \<Longrightarrow>
+lemma dvd_multiplicity_eq_nat: "0 < (x::nat) \<Longrightarrow> 0 < y \<Longrightarrow>
(x dvd y) = (ALL p. multiplicity p x <= multiplicity p y)"
- by (auto intro: nat_dvd_multiplicity nat_multiplicity_dvd)
+ by (auto intro: dvd_multiplicity_nat multiplicity_dvd_nat)
-lemma int_dvd_multiplicity_eq: "0 < (x::int) \<Longrightarrow> 0 < y \<Longrightarrow>
+lemma dvd_multiplicity_eq_int: "0 < (x::int) \<Longrightarrow> 0 < y \<Longrightarrow>
(x dvd y) = (ALL p >= 0. multiplicity p x <= multiplicity p y)"
- by (auto intro: int_dvd_multiplicity int_multiplicity_dvd)
+ by (auto intro: dvd_multiplicity_int multiplicity_dvd_int)
-lemma nat_prime_factors_altdef2: "(n::nat) > 0 \<Longrightarrow>
+lemma prime_factors_altdef2_nat: "(n::nat) > 0 \<Longrightarrow>
(p : prime_factors n) = (prime p & p dvd n)"
apply (case_tac "prime p")
apply auto
- apply (subst nat_prime_factorization [where n = n], assumption)
+ apply (subst prime_factorization_nat [where n = n], assumption)
apply (rule dvd_trans)
apply (rule dvd_power [where x = p and n = "multiplicity p n"])
- apply (subst (asm) nat_prime_factors_altdef, force)
+ apply (subst (asm) prime_factors_altdef_nat, force)
apply (rule dvd_setprod)
apply auto
- apply (subst nat_prime_factors_altdef)
- apply (subst (asm) nat_dvd_multiplicity_eq)
+ apply (subst prime_factors_altdef_nat)
+ apply (subst (asm) dvd_multiplicity_eq_nat)
apply auto
apply (drule spec [where x = p])
apply auto
done
-lemma int_prime_factors_altdef2:
+lemma prime_factors_altdef2_int:
assumes "(n::int) > 0"
shows "(p : prime_factors n) = (prime p & p dvd n)"
apply (case_tac "p >= 0")
- apply (rule nat_prime_factors_altdef2 [transferred])
+ apply (rule prime_factors_altdef2_nat [transferred])
using prems apply auto
- apply (auto simp add: int_prime_ge_0 int_prime_factors_ge_0)
+ apply (auto simp add: prime_ge_0_int prime_factors_ge_0_int)
done
-lemma nat_multiplicity_eq:
+lemma multiplicity_eq_nat:
fixes x and y::nat
assumes [arith]: "x > 0" "y > 0" and
mult_eq [simp]: "!!p. prime p \<Longrightarrow> multiplicity p x = multiplicity p y"
shows "x = y"
apply (rule dvd_anti_sym)
- apply (auto intro: nat_multiplicity_dvd')
+ apply (auto intro: multiplicity_dvd'_nat)
done
-lemma int_multiplicity_eq:
+lemma multiplicity_eq_int:
fixes x and y::int
assumes [arith]: "x > 0" "y > 0" and
mult_eq [simp]: "!!p. prime p \<Longrightarrow> multiplicity p x = multiplicity p y"
shows "x = y"
apply (rule dvd_anti_sym [transferred])
- apply (auto intro: int_multiplicity_dvd')
+ apply (auto intro: multiplicity_dvd'_int)
done
subsection {* An application *}
-lemma nat_gcd_eq:
+lemma gcd_eq_nat:
assumes pos [arith]: "x > 0" "y > 0"
shows "gcd (x::nat) y =
(PROD p: prime_factors x Un prime_factors y.
@@ -874,26 +874,26 @@
have aux: "!!p. prime p \<Longrightarrow> multiplicity p z =
min (multiplicity p x) (multiplicity p y)"
unfolding z_def
- apply (subst nat_multiplicity_prod_prime_powers)
- apply (auto simp add: nat_multiplicity_not_factor)
+ apply (subst multiplicity_prod_prime_powers_nat)
+ apply (auto simp add: multiplicity_not_factor_nat)
done
have "z dvd x"
- by (intro nat_multiplicity_dvd', auto simp add: aux)
+ by (intro multiplicity_dvd'_nat, auto simp add: aux)
moreover have "z dvd y"
- by (intro nat_multiplicity_dvd', auto simp add: aux)
+ by (intro multiplicity_dvd'_nat, auto simp add: aux)
moreover have "ALL w. w dvd x & w dvd y \<longrightarrow> w dvd z"
apply auto
apply (case_tac "w = 0", auto)
- apply (erule nat_multiplicity_dvd')
- apply (auto intro: nat_dvd_multiplicity simp add: aux)
+ apply (erule multiplicity_dvd'_nat)
+ apply (auto intro: dvd_multiplicity_nat simp add: aux)
done
ultimately have "z = gcd x y"
- by (subst nat_gcd_unique [symmetric], blast)
+ by (subst gcd_unique_nat [symmetric], blast)
thus ?thesis
unfolding z_def by auto
qed
-lemma nat_lcm_eq:
+lemma lcm_eq_nat:
assumes pos [arith]: "x > 0" "y > 0"
shows "lcm (x::nat) y =
(PROD p: prime_factors x Un prime_factors y.
@@ -906,61 +906,61 @@
have aux: "!!p. prime p \<Longrightarrow> multiplicity p z =
max (multiplicity p x) (multiplicity p y)"
unfolding z_def
- apply (subst nat_multiplicity_prod_prime_powers)
- apply (auto simp add: nat_multiplicity_not_factor)
+ apply (subst multiplicity_prod_prime_powers_nat)
+ apply (auto simp add: multiplicity_not_factor_nat)
done
have "x dvd z"
- by (intro nat_multiplicity_dvd', auto simp add: aux)
+ by (intro multiplicity_dvd'_nat, auto simp add: aux)
moreover have "y dvd z"
- by (intro nat_multiplicity_dvd', auto simp add: aux)
+ by (intro multiplicity_dvd'_nat, auto simp add: aux)
moreover have "ALL w. x dvd w & y dvd w \<longrightarrow> z dvd w"
apply auto
apply (case_tac "w = 0", auto)
- apply (rule nat_multiplicity_dvd')
- apply (auto intro: nat_dvd_multiplicity simp add: aux)
+ apply (rule multiplicity_dvd'_nat)
+ apply (auto intro: dvd_multiplicity_nat simp add: aux)
done
ultimately have "z = lcm x y"
- by (subst nat_lcm_unique [symmetric], blast)
+ by (subst lcm_unique_nat [symmetric], blast)
thus ?thesis
unfolding z_def by auto
qed
-lemma nat_multiplicity_gcd:
+lemma multiplicity_gcd_nat:
assumes [arith]: "x > 0" "y > 0"
shows "multiplicity (p::nat) (gcd x y) =
min (multiplicity p x) (multiplicity p y)"
- apply (subst nat_gcd_eq)
+ apply (subst gcd_eq_nat)
apply auto
- apply (subst nat_multiplicity_prod_prime_powers)
+ apply (subst multiplicity_prod_prime_powers_nat)
apply auto
done
-lemma nat_multiplicity_lcm:
+lemma multiplicity_lcm_nat:
assumes [arith]: "x > 0" "y > 0"
shows "multiplicity (p::nat) (lcm x y) =
max (multiplicity p x) (multiplicity p y)"
- apply (subst nat_lcm_eq)
+ apply (subst lcm_eq_nat)
apply auto
- apply (subst nat_multiplicity_prod_prime_powers)
+ apply (subst multiplicity_prod_prime_powers_nat)
apply auto
done
-lemma nat_gcd_lcm_distrib: "gcd (x::nat) (lcm y z) = lcm (gcd x y) (gcd x z)"
+lemma gcd_lcm_distrib_nat: "gcd (x::nat) (lcm y z) = lcm (gcd x y) (gcd x z)"
apply (case_tac "x = 0 | y = 0 | z = 0")
apply auto
- apply (rule nat_multiplicity_eq)
- apply (auto simp add: nat_multiplicity_gcd nat_multiplicity_lcm
- nat_lcm_pos)
+ apply (rule multiplicity_eq_nat)
+ apply (auto simp add: multiplicity_gcd_nat multiplicity_lcm_nat
+ lcm_pos_nat)
done
-lemma int_gcd_lcm_distrib: "gcd (x::int) (lcm y z) = lcm (gcd x y) (gcd x z)"
- apply (subst (1 2 3) int_gcd_abs)
- apply (subst int_lcm_abs)
+lemma gcd_lcm_distrib_int: "gcd (x::int) (lcm y z) = lcm (gcd x y) (gcd x z)"
+ apply (subst (1 2 3) gcd_abs_int)
+ apply (subst lcm_abs_int)
apply (subst (2) abs_of_nonneg)
apply force
- apply (rule nat_gcd_lcm_distrib [transferred])
+ apply (rule gcd_lcm_distrib_nat [transferred])
apply auto
done
--- a/src/HOL/RealDef.thy Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/RealDef.thy Thu Jul 09 08:55:42 2009 +0200
@@ -925,7 +925,7 @@
have "?gcd' = 1"
proof -
have "?gcd * ?gcd' = gcd (?gcd * ?k) (?gcd * ?l)"
- by (rule nat_gcd_mult_distrib)
+ by (rule gcd_mult_distrib_nat)
with gcd_k gcd_l have "?gcd * ?gcd' = ?gcd" by simp
with gcd show ?thesis by auto
qed
--- a/src/HOL/Set.thy Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/Set.thy Thu Jul 09 08:55:42 2009 +0200
@@ -449,7 +449,7 @@
lemma ballE [elim]: "ALL x:A. P x ==> (P x ==> Q) ==> (x ~: A ==> Q) ==> Q"
by (unfold Ball_def) blast
-ML {* bind_thm ("rev_ballE", permute_prems 1 1 @{thm ballE}) *}
+ML {* bind_thm ("rev_ballE", Thm.permute_prems 1 1 @{thm ballE}) *}
text {*
\medskip This tactic takes assumptions @{prop "ALL x:A. P x"} and
--- a/src/HOL/TLA/Intensional.thy Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/TLA/Intensional.thy Thu Jul 09 08:55:42 2009 +0200
@@ -268,7 +268,7 @@
let
(* analogous to RS, but using matching instead of resolution *)
fun matchres tha i thb =
- case Seq.chop 2 (biresolution true [(false,tha)] i thb) of
+ case Seq.chop 2 (Thm.biresolution true [(false,tha)] i thb) of
([th],_) => th
| ([],_) => raise THM("matchres: no match", i, [tha,thb])
| _ => raise THM("matchres: multiple unifiers", i, [tha,thb])
--- a/src/HOL/Tools/Datatype/datatype_rep_proofs.ML Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/Tools/Datatype/datatype_rep_proofs.ML Thu Jul 09 08:55:42 2009 +0200
@@ -37,10 +37,6 @@
(** theory context references **)
-val f_myinv_f = thm "f_myinv_f";
-val myinv_f_f = thm "myinv_f_f";
-
-
fun exh_thm_of (dt_info : info Symtab.table) tname =
#exhaustion (the (Symtab.lookup dt_info tname));
@@ -162,7 +158,7 @@
app_bnds (mk_Free "x" (Ts ---> Univ_elT) j) (length Ts)
in (j + 1, list_all (map (pair "x") Ts,
HOLogic.mk_Trueprop
- (Free (List.nth (rep_set_names', k), UnivT') $ free_t)) :: prems,
+ (Free (nth rep_set_names' k, UnivT') $ free_t)) :: prems,
mk_lim free_t Ts :: ts)
end
| _ =>
@@ -225,7 +221,7 @@
val free_t = mk_Free "x" T j
in (case (strip_dtyp dt, strip_type T) of
((_, DtRec m), (Us, U)) => (j + 1, free_t :: l_args, mk_lim
- (Const (List.nth (all_rep_names, m), U --> Univ_elT) $
+ (Const (nth all_rep_names m, U --> Univ_elT) $
app_bnds free_t (length Us)) Us :: r_args)
| _ => (j + 1, free_t::l_args, (Leaf $ mk_inj T free_t)::r_args))
end;
@@ -295,8 +291,8 @@
fun make_iso_def k ks n ((fs, eqns, i), (cname, cargs)) =
let
val argTs = map (typ_of_dtyp descr' sorts) cargs;
- val T = List.nth (recTs, k);
- val rep_name = List.nth (all_rep_names, k);
+ val T = nth recTs k;
+ val rep_name = nth all_rep_names k;
val rep_const = Const (rep_name, T --> Univ_elT);
val constr = Const (cname, argTs ---> T);
@@ -311,7 +307,7 @@
Ts @ [Us ---> Univ_elT])
else
(i2 + 1, i2', ts @ [mk_lim
- (Const (List.nth (all_rep_names, j), U --> Univ_elT) $
+ (Const (nth all_rep_names j, U --> Univ_elT) $
app_bnds (mk_Free "x" T' i2) (length Us)) Us], Ts)
| _ => (i2 + 1, i2', ts @ [Leaf $ mk_inj T' (mk_Free "x" T' i2)], Ts))
end;
@@ -339,7 +335,7 @@
let
val (fs', eqns', _) = Library.foldl (make_iso_def k ks (length constrs))
((fs, eqns, 1), constrs);
- val iso = (List.nth (recTs, k), List.nth (all_rep_names, k))
+ val iso = (nth recTs k, nth all_rep_names k)
in (fs', eqns', isos @ [iso]) end;
val (fs, eqns, isos) = Library.foldl process_dt (([], [], []), ds);
@@ -397,9 +393,9 @@
fun mk_ind_concl (i, _) =
let
- val T = List.nth (recTs, i);
- val Rep_t = Const (List.nth (all_rep_names, i), T --> Univ_elT);
- val rep_set_name = List.nth (rep_set_names, i)
+ val T = nth recTs i;
+ val Rep_t = Const (nth all_rep_names i, T --> Univ_elT);
+ val rep_set_name = nth rep_set_names i
in (HOLogic.all_const T $ Abs ("y", T, HOLogic.imp $
HOLogic.mk_eq (Rep_t $ mk_Free "x" T i, Rep_t $ Bound 0) $
HOLogic.mk_eq (mk_Free "x" T i, Bound 0)),
@@ -490,7 +486,7 @@
val Abs_inverse_thms' =
map #1 newT_iso_axms @
- map2 (fn r_inj => fn r => f_myinv_f OF [r_inj, r RS mp])
+ map2 (fn r_inj => fn r => @{thm f_inv_f} OF [r_inj, r RS mp])
iso_inj_thms_unfolded iso_thms;
val Abs_inverse_thms = maps (mk_funs_inv thy5) Abs_inverse_thms';
@@ -578,22 +574,22 @@
val _ = message config "Proving induction rule for datatypes ...";
val Rep_inverse_thms = (map (fn (_, iso, _) => iso RS subst) newT_iso_axms) @
- (map (fn r => r RS myinv_f_f RS subst) iso_inj_thms_unfolded);
- val Rep_inverse_thms' = map (fn r => r RS myinv_f_f) iso_inj_thms_unfolded;
+ (map (fn r => r RS @{thm inv_f_f} RS subst) iso_inj_thms_unfolded);
+ val Rep_inverse_thms' = map (fn r => r RS @{thm inv_f_f}) iso_inj_thms_unfolded;
fun mk_indrule_lemma ((prems, concls), ((i, _), T)) =
let
- val Rep_t = Const (List.nth (all_rep_names, i), T --> Univ_elT) $
+ val Rep_t = Const (nth all_rep_names i, T --> Univ_elT) $
mk_Free "x" T i;
val Abs_t = if i < length newTs then
Const (Sign.intern_const thy6
- ("Abs_" ^ (List.nth (new_type_names, i))), Univ_elT --> T)
- else Const ("Inductive.myinv", [T --> Univ_elT, Univ_elT] ---> T) $
- Const (List.nth (all_rep_names, i), T --> Univ_elT)
+ ("Abs_" ^ (nth new_type_names i)), Univ_elT --> T)
+ else Const (@{const_name Fun.inv}, [T --> Univ_elT, Univ_elT] ---> T) $
+ Const (nth all_rep_names i, T --> Univ_elT)
in (prems @ [HOLogic.imp $
- (Const (List.nth (rep_set_names, i), UnivT') $ Rep_t) $
+ (Const (nth rep_set_names i, UnivT') $ Rep_t) $
(mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t))],
concls @ [mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ mk_Free "x" T i])
end;
--- a/src/HOL/Tools/TFL/rules.ML Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/Tools/TFL/rules.ML Thu Jul 09 08:55:42 2009 +0200
@@ -519,14 +519,14 @@
handle TERM _ => get (rst, n+1, L)
| U.ERR _ => get (rst, n+1, L);
-(* Note: rename_params_rule counts from 1, not 0 *)
+(* Note: Thm.rename_params_rule counts from 1, not 0 *)
fun rename thm =
let val thy = Thm.theory_of_thm thm
val tych = cterm_of thy
val ants = Logic.strip_imp_prems (Thm.prop_of thm)
val news = get (ants,1,[])
in
- fold rename_params_rule news thm
+ fold Thm.rename_params_rule news thm
end;
--- a/src/HOL/Tools/meson.ML Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/Tools/meson.ML Thu Jul 09 08:55:42 2009 +0200
@@ -470,7 +470,7 @@
(*Like TRYALL eq_assume_tac, but avoids expensive THEN calls*)
fun TRYING_eq_assume_tac 0 st = Seq.single st
| TRYING_eq_assume_tac i st =
- TRYING_eq_assume_tac (i-1) (eq_assumption i st)
+ TRYING_eq_assume_tac (i-1) (Thm.eq_assumption i st)
handle THM _ => TRYING_eq_assume_tac (i-1) st;
fun TRYALL_eq_assume_tac st = TRYING_eq_assume_tac (nprems_of st) st;
--- a/src/HOL/Tools/metis_tools.ML Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/Tools/metis_tools.ML Thu Jul 09 08:55:42 2009 +0200
@@ -339,7 +339,7 @@
could then fail. See comment on mk_var.*)
fun resolve_inc_tyvars(tha,i,thb) =
let val tha = Drule.incr_type_indexes (1 + Thm.maxidx_of thb) tha
- val ths = Seq.list_of (bicompose false (false,tha,nprems_of tha) i thb)
+ val ths = Seq.list_of (Thm.bicompose false (false,tha,nprems_of tha) i thb)
in
case distinct Thm.eq_thm ths of
[th] => th
--- a/src/HOL/Tools/quickcheck_generators.ML Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/Tools/quickcheck_generators.ML Thu Jul 09 08:55:42 2009 +0200
@@ -5,7 +5,6 @@
signature QUICKCHECK_GENERATORS =
sig
- val compile_generator_expr: theory -> term -> int -> term list option
type seed = Random_Engine.seed
val random_fun: typ -> typ -> ('a -> 'a -> bool) -> ('a -> term)
-> (seed -> ('b * (unit -> term)) * seed) -> (seed -> seed * seed)
@@ -16,6 +15,7 @@
-> string list -> string list * string list -> typ list * typ list
-> term list * (term * term) list
val ensure_random_datatype: Datatype.config -> string list -> theory -> theory
+ val compile_generator_expr: theory -> term -> int -> term list option
val eval_ref: (unit -> int -> seed -> term list option * seed) option ref
val setup: theory -> theory
end;
@@ -23,42 +23,13 @@
structure Quickcheck_Generators : QUICKCHECK_GENERATORS =
struct
-(** building and compiling generator expressions **)
-
-val eval_ref : (unit -> int -> int * int -> term list option * (int * int)) option ref = ref NONE;
-
-val target = "Quickcheck";
+(** abstract syntax **)
-fun mk_generator_expr thy prop tys =
- let
- val bound_max = length tys - 1;
- val bounds = map_index (fn (i, ty) =>
- (2 * (bound_max - i) + 1, 2 * (bound_max - i), 2 * i, ty)) tys;
- val result = list_comb (prop, map (fn (i, _, _, _) => Bound i) bounds);
- val terms = HOLogic.mk_list @{typ term} (map (fn (_, i, _, _) => Bound i $ @{term "()"}) bounds);
- val check = @{term "If :: bool => term list option => term list option => term list option"}
- $ result $ @{term "None :: term list option"} $ (@{term "Some :: term list => term list option "} $ terms);
- val return = @{term "Pair :: term list option => Random.seed => term list option * Random.seed"};
- fun liftT T sT = sT --> HOLogic.mk_prodT (T, sT);
- fun mk_termtyp ty = HOLogic.mk_prodT (ty, @{typ "unit => term"});
- fun mk_scomp T1 T2 sT f g = Const (@{const_name scomp},
- liftT T1 sT --> (T1 --> liftT T2 sT) --> liftT T2 sT) $ f $ g;
- fun mk_split ty = Sign.mk_const thy
- (@{const_name split}, [ty, @{typ "unit => term"}, liftT @{typ "term list option"} @{typ Random.seed}]);
- fun mk_scomp_split ty t t' =
- mk_scomp (mk_termtyp ty) @{typ "term list option"} @{typ Random.seed} t
- (mk_split ty $ Abs ("", ty, Abs ("", @{typ "unit => term"}, t')));
- fun mk_bindclause (_, _, i, ty) = mk_scomp_split ty
- (Sign.mk_const thy (@{const_name Quickcheck.random}, [ty]) $ Bound i);
- in Abs ("n", @{typ code_numeral}, fold_rev mk_bindclause bounds (return $ check)) end;
-
-fun compile_generator_expr thy t =
- let
- val tys = (map snd o fst o strip_abs) t;
- val t' = mk_generator_expr thy t tys;
- val f = Code_ML.eval (SOME target) ("Quickcheck_Generators.eval_ref", eval_ref)
- (fn proc => fn g => fn s => g s #>> (Option.map o map) proc) thy t' [];
- in f #> Random_Engine.run end;
+fun termifyT T = HOLogic.mk_prodT (T, @{typ "unit => term"})
+val size = @{term "i::code_numeral"};
+val size1 = @{term "(i::code_numeral) - 1"};
+val size' = @{term "j::code_numeral"};
+val seed = @{term "s::Random.seed"};
(** typ "'a => 'b" **)
@@ -91,25 +62,22 @@
(** type copies **)
-fun mk_random_typecopy tyco vs constr typ thy =
+fun mk_random_typecopy tyco vs constr T' thy =
let
val Ts = map TFree vs;
val T = Type (tyco, Ts);
- fun mk_termifyT T = HOLogic.mk_prodT (T, @{typ "unit => term"})
- val Ttm = mk_termifyT T;
- val typtm = mk_termifyT typ;
+ val Tm = termifyT T;
+ val Tm' = termifyT T';
fun mk_const c Ts = Const (c, Sign.const_instance thy (c, Ts));
- fun mk_random T = mk_const @{const_name Quickcheck.random} [T];
- val size = @{term "j::code_numeral"};
val v = "x";
- val t_v = Free (v, typtm);
+ val t_v = Free (v, Tm');
val t_constr = mk_const constr Ts;
- val lhs = mk_random T $ size;
- val rhs = HOLogic.mk_ST [(((mk_random typ) $ size, @{typ Random.seed}), SOME (v, typtm))]
- (HOLogic.mk_return Ttm @{typ Random.seed}
- (mk_const "Code_Eval.valapp" [typ, T]
+ val lhs = HOLogic.mk_random T size;
+ val rhs = HOLogic.mk_ST [((HOLogic.mk_random T' size, @{typ Random.seed}), SOME (v, Tm'))]
+ (HOLogic.mk_return Tm @{typ Random.seed}
+ (mk_const "Code_Eval.valapp" [T', T]
$ HOLogic.mk_prod (t_constr, Abs ("u", @{typ unit}, HOLogic.reflect_term t_constr)) $ t_v))
- @{typ Random.seed} (SOME Ttm, @{typ Random.seed});
+ @{typ Random.seed} (SOME Tm, @{typ Random.seed});
val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));
in
thy
@@ -122,16 +90,16 @@
fun ensure_random_typecopy tyco thy =
let
- val SOME { vs = raw_vs, constr, typ = raw_typ, ... } =
+ val SOME { vs = raw_vs, constr, typ = raw_T, ... } =
Typecopy.get_info thy tyco;
val constrain = curry (Sorts.inter_sort (Sign.classes_of thy));
- val typ = map_atyps (fn TFree (v, sort) =>
- TFree (v, constrain sort @{sort random})) raw_typ;
- val vs' = Term.add_tfreesT typ [];
+ val T = map_atyps (fn TFree (v, sort) =>
+ TFree (v, constrain sort @{sort random})) raw_T;
+ val vs' = Term.add_tfreesT T [];
val vs = map (fn (v, sort) =>
(v, the_default (constrain sort @{sort typerep}) (AList.lookup (op =) vs' v))) raw_vs;
- val do_inst = Sign.of_sort thy (typ, @{sort random});
- in if do_inst then mk_random_typecopy tyco vs constr typ thy else thy end;
+ val can_inst = Sign.of_sort thy (T, @{sort random});
+ in if can_inst then mk_random_typecopy tyco vs constr T thy else thy end;
(** datatypes **)
@@ -258,12 +226,7 @@
fun mk_random_aux_eqs thy descr vs tycos (names, auxnames) (Ts, Us) =
let
val mk_const = curry (Sign.mk_const thy);
- val i = @{term "i\<Colon>code_numeral"};
- val i1 = @{term "(i\<Colon>code_numeral) - 1"};
- val j = @{term "j\<Colon>code_numeral"};
- val seed = @{term "s\<Colon>Random.seed"};
val random_auxsN = map (prefix (random_auxN ^ "_")) (names @ auxnames);
- fun termifyT T = HOLogic.mk_prodT (T, @{typ "unit \<Rightarrow> term"});
val rTs = Ts @ Us;
fun random_resultT T = @{typ Random.seed}
--> HOLogic.mk_prodT (termifyT T,@{typ Random.seed});
@@ -272,7 +235,7 @@
val random_auxT = sizeT o random_resultT;
val random_auxs = map2 (fn s => fn rT => Free (s, random_auxT rT))
random_auxsN rTs;
- fun mk_random_call T = (NONE, (HOLogic.mk_random T j, T));
+ fun mk_random_call T = (NONE, (HOLogic.mk_random T size', T));
fun mk_random_aux_call fTs (k, _) (tyco, Ts) =
let
val T = Type (tyco, Ts);
@@ -280,7 +243,7 @@
| mk_random_fun_lift (fT :: fTs) t =
mk_const @{const_name random_fun_lift} [fTs ---> T, fT] $
mk_random_fun_lift fTs t;
- val t = mk_random_fun_lift fTs (nth random_auxs k $ i1 $ j);
+ val t = mk_random_fun_lift fTs (nth random_auxs k $ size1 $ size');
val size = Option.map snd (DatatypeCodegen.find_shortest_path descr k)
|> the_default 0;
in (SOME size, (t, fTs ---> T)) end;
@@ -300,9 +263,9 @@
val t = HOLogic.mk_ST (map (fn (t, _) => (t, @{typ Random.seed})) tTs ~~ map SOME vs')
tc @{typ Random.seed} (SOME T, @{typ Random.seed});
val tk = if is_rec
- then if k = 0 then i
+ then if k = 0 then size
else @{term "Quickcheck.beyond :: code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"}
- $ HOLogic.mk_number @{typ code_numeral} k $ i
+ $ HOLogic.mk_number @{typ code_numeral} k $ size
else @{term "1::code_numeral"}
in (is_rec, HOLogic.mk_prod (tk, t)) end;
fun sort_rec xs =
@@ -316,24 +279,23 @@
$ (mk_const @{const_name Random.select_weight} [random_resultT rT]
$ HOLogic.mk_list (HOLogic.mk_prodT (@{typ code_numeral}, random_resultT rT)) xs)
$ seed;
- val auxs_lhss = map (fn t => t $ i $ j $ seed) random_auxs;
+ val auxs_lhss = map (fn t => t $ size $ size' $ seed) random_auxs;
val auxs_rhss = map mk_select gen_exprss;
in (random_auxs, auxs_lhss ~~ auxs_rhss) end;
fun mk_random_datatype config descr vs tycos prfx (names, auxnames) (Ts, Us) thy =
let
val _ = DatatypeAux.message config "Creating quickcheck generators ...";
- val i = @{term "i\<Colon>code_numeral"};
val mk_prop_eq = HOLogic.mk_Trueprop o HOLogic.mk_eq;
fun mk_size_arg k = case DatatypeCodegen.find_shortest_path descr k
- of SOME (_, l) => if l = 0 then i
+ of SOME (_, l) => if l = 0 then size
else @{term "max :: code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"}
- $ HOLogic.mk_number @{typ code_numeral} l $ i
- | NONE => i;
+ $ HOLogic.mk_number @{typ code_numeral} l $ size
+ | NONE => size;
val (random_auxs, auxs_eqs) = (apsnd o map) mk_prop_eq
(mk_random_aux_eqs thy descr vs tycos (names, auxnames) (Ts, Us));
val random_defs = map_index (fn (k, T) => mk_prop_eq
- (HOLogic.mk_random T i, nth random_auxs k $ mk_size_arg k $ i)) Ts;
+ (HOLogic.mk_random T size, nth random_auxs k $ mk_size_arg k $ size)) Ts;
in
thy
|> TheoryTarget.instantiation (tycos, vs, @{sort random})
@@ -381,11 +343,49 @@
end;
+(** building and compiling generator expressions **)
+
+val eval_ref : (unit -> int -> int * int -> term list option * (int * int)) option ref = ref NONE;
+
+val target = "Quickcheck";
+
+fun mk_generator_expr thy prop Ts =
+ let
+ val bound_max = length Ts - 1;
+ val bounds = map_index (fn (i, ty) =>
+ (2 * (bound_max - i) + 1, 2 * (bound_max - i), 2 * i, ty)) Ts;
+ val result = list_comb (prop, map (fn (i, _, _, _) => Bound i) bounds);
+ val terms = HOLogic.mk_list @{typ term} (map (fn (_, i, _, _) => Bound i $ @{term "()"}) bounds);
+ val check = @{term "If :: bool => term list option => term list option => term list option"}
+ $ result $ @{term "None :: term list option"} $ (@{term "Some :: term list => term list option "} $ terms);
+ val return = @{term "Pair :: term list option => Random.seed => term list option * Random.seed"};
+ fun liftT T sT = sT --> HOLogic.mk_prodT (T, sT);
+ fun mk_termtyp T = HOLogic.mk_prodT (T, @{typ "unit => term"});
+ fun mk_scomp T1 T2 sT f g = Const (@{const_name scomp},
+ liftT T1 sT --> (T1 --> liftT T2 sT) --> liftT T2 sT) $ f $ g;
+ fun mk_split T = Sign.mk_const thy
+ (@{const_name split}, [T, @{typ "unit => term"}, liftT @{typ "term list option"} @{typ Random.seed}]);
+ fun mk_scomp_split T t t' =
+ mk_scomp (mk_termtyp T) @{typ "term list option"} @{typ Random.seed} t
+ (mk_split T $ Abs ("", T, Abs ("", @{typ "unit => term"}, t')));
+ fun mk_bindclause (_, _, i, T) = mk_scomp_split T
+ (Sign.mk_const thy (@{const_name Quickcheck.random}, [T]) $ Bound i);
+ in Abs ("n", @{typ code_numeral}, fold_rev mk_bindclause bounds (return $ check)) end;
+
+fun compile_generator_expr thy t =
+ let
+ val Ts = (map snd o fst o strip_abs) t;
+ val t' = mk_generator_expr thy t Ts;
+ val compile = Code_ML.eval (SOME target) ("Quickcheck_Generators.eval_ref", eval_ref)
+ (fn proc => fn g => fn s => g s #>> (Option.map o map) proc) thy t' [];
+ in compile #> Random_Engine.run end;
+
+
(** setup **)
-val setup = Code_Target.extend_target (target, (Code_ML.target_Eval, K I))
- #> Quickcheck.add_generator ("code", compile_generator_expr o ProofContext.theory_of)
- #> Typecopy.interpretation ensure_random_typecopy
- #> Datatype.interpretation ensure_random_datatype;
+val setup = Typecopy.interpretation ensure_random_typecopy
+ #> Datatype.interpretation ensure_random_datatype
+ #> Code_Target.extend_target (target, (Code_ML.target_Eval, K I))
+ #> Quickcheck.add_generator ("code", compile_generator_expr o ProofContext.theory_of);
end;
--- a/src/HOL/Tools/recfun_codegen.ML Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/Tools/recfun_codegen.ML Thu Jul 09 08:55:42 2009 +0200
@@ -26,25 +26,17 @@
fun add_thm NONE thm thy = Code.add_eqn thm thy
| add_thm (SOME module_name) thm thy =
let
- val (thm', _) = Code.mk_eqn thy (K false) (thm, true)
+ val (thm', _) = Code.mk_eqn thy (thm, true)
in
thy
- |> ModuleData.map (Symtab.update (Code.const_eqn thy thm', module_name))
+ |> ModuleData.map (Symtab.update (fst (Code.const_typ_eqn thy thm'), module_name))
|> Code.add_eqn thm'
end;
-fun meta_eq_to_obj_eq thy thm =
- let
- val T = (fastype_of o fst o Logic.dest_equals o Thm.prop_of) thm;
- in if Sign.of_sort thy (T, @{sort type})
- then SOME (Conv.fconv_rule Drule.beta_eta_conversion (@{thm meta_eq_to_obj_eq} OF [thm]))
- else NONE
- end;
-
fun expand_eta thy [] = []
| expand_eta thy (thms as thm :: _) =
let
- val (_, ty) = Code.const_typ_eqn thm;
+ val (_, ty) = Code.const_typ_eqn thy thm;
in if null (Term.add_tvarsT ty []) orelse (null o fst o strip_type) ty
then thms
else map (Code.expand_eta thy 1) thms
@@ -57,12 +49,11 @@
val thms = Code.these_eqns thy c'
|> map_filter (fn (thm, linear) => if linear then SOME thm else NONE)
|> expand_eta thy
- |> map_filter (meta_eq_to_obj_eq thy)
|> Code.norm_varnames thy
|> map (rpair opt_name)
in if null thms then NONE else SOME thms end;
-val dest_eqn = HOLogic.dest_eq o HOLogic.dest_Trueprop;
+val dest_eqn = Logic.dest_equals;
val const_of = dest_Const o head_of o fst o dest_eqn;
fun get_equations thy defs (s, T) =
--- a/src/HOL/Tools/refute.ML Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/Tools/refute.ML Thu Jul 09 08:55:42 2009 +0200
@@ -909,8 +909,8 @@
(* and the class definition *)
let
val class = Logic.class_of_const s
- val inclass = Logic.mk_inclass (TVar (("'a", 0), [class]), class)
- val ax_in = SOME (specialize_type thy (s, T) inclass)
+ val of_class = Logic.mk_of_class (TVar (("'a", 0), [class]), class)
+ val ax_in = SOME (specialize_type thy (s, T) of_class)
(* type match may fail due to sort constraints *)
handle Type.TYPE_MATCH => NONE
val ax_1 = Option.map (fn ax => (Syntax.string_of_term_global thy ax, ax))
--- a/src/HOL/ex/LocaleTest2.thy Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/ex/LocaleTest2.thy Thu Jul 09 08:55:42 2009 +0200
@@ -599,7 +599,7 @@
apply (rule_tac x = "gcd x y" in exI)
apply auto [1]
apply (rule_tac x = "lcm x y" in exI)
- apply (auto intro: nat_lcm_dvd1 nat_lcm_dvd2 nat_lcm_least)
+ apply (auto intro: lcm_dvd1_nat lcm_dvd2_nat lcm_least_nat)
done
then interpret nat_dvd: dlat "op dvd :: [nat, nat] => bool" .
txt {* Interpretation to ease use of definitions, which are
@@ -613,7 +613,7 @@
apply (unfold nat_dvd.join_def)
apply (rule the_equality)
apply (unfold nat_dvd.is_sup_def)
- by (auto intro: nat_lcm_dvd1 nat_lcm_dvd2 nat_lcm_least)
+ by (auto intro: lcm_dvd1_nat lcm_dvd2_nat lcm_least_nat)
qed
text {* Interpreted theorems from the locales, involving defined terms. *}
--- a/src/HOL/ex/Sqrt.thy Sun Jul 05 21:10:23 2009 +0200
+++ b/src/HOL/ex/Sqrt.thy Thu Jul 09 08:55:42 2009 +0200
@@ -34,12 +34,12 @@
have "p dvd m \<and> p dvd n"
proof
from eq have "p dvd m\<twosuperior>" ..
- with `prime p` pos2 show "p dvd m" by (rule nat_prime_dvd_power)
+ with `prime p` pos2 show "p dvd m" by (rule prime_dvd_power_nat)
then obtain k where "m = p * k" ..
with eq have "p * n\<twosuperior> = p\<twosuperior> * k\<twosuperior>" by (auto simp add: power2_eq_square mult_ac)
with p have "n\<twosuperior> = p * k\<twosuperior>" by (simp add: power2_eq_square)
then have "p dvd n\<twosuperior>" ..
- with `prime p` pos2 show "p dvd n" by (rule nat_prime_dvd_power)
+ with `prime p` pos2 show "p dvd n" by (rule prime_dvd_power_nat)
qed
then have "p dvd gcd m n" ..
with gcd have "p dvd 1" by simp
@@ -48,7 +48,7 @@
qed
corollary "sqrt (real (2::nat)) \<notin> \<rat>"
- by (rule sqrt_prime_irrational) (rule nat_two_is_prime)
+ by (rule sqrt_prime_irrational) (rule two_is_prime_nat)
subsection {* Variations *}
@@ -75,13 +75,13 @@
also have "\<dots> * real (n\<twosuperior>) = real (p * n\<twosuperior>)" by simp
finally have eq: "m\<twosuperior> = p * n\<twosuperior>" ..
then have "p dvd m\<twosuperior>" ..
- with `prime p` pos2 have dvd_m: "p dvd m" by (rule nat_prime_dvd_power)
+ with `prime p` pos2 have dvd_m: "p dvd m" by (rule prime_dvd_power_nat)
then obtain k where "m = p * k" ..
with eq have "p * n\<twosuperior> = p\<twosuperior> * k\<twosuperior>" by (auto simp add: power2_eq_square mult_ac)
with p have "n\<twosuperior> = p * k\<twosuperior>" by (simp add: power2_eq_square)
then have "p dvd n\<twosuperior>" ..
- with `prime p` pos2 have "p dvd n" by (rule nat_prime_dvd_power)
- with dvd_m have "p dvd gcd m n" by (rule nat_gcd_greatest)
+ with `prime p` pos2 have "p dvd n" by (rule prime_dvd_power_nat)
+ with dvd_m have "p dvd gcd m n" by (rule gcd_greatest_nat)
with gcd have "p dvd 1" by simp
then have "p \<le> 1" by (simp add: dvd_imp_le)
with p show False by simp
--- a/src/Provers/blast.ML Sun Jul 05 21:10:23 2009 +0200
+++ b/src/Provers/blast.ML Thu Jul 09 08:55:42 2009 +0200
@@ -473,7 +473,7 @@
(*Like dup_elim, but puts the duplicated major premise FIRST*)
-fun rev_dup_elim th = (th RSN (2, revcut_rl)) |> assumption 2 |> Seq.hd;
+fun rev_dup_elim th = (th RSN (2, revcut_rl)) |> Thm.assumption 2 |> Seq.hd;
(*Rotate the assumptions in all new subgoals for the LIFO discipline*)
@@ -485,7 +485,7 @@
fun rot_tac [] i st = Seq.single st
| rot_tac (0::ks) i st = rot_tac ks (i+1) st
- | rot_tac (k::ks) i st = rot_tac ks (i+1) (rotate_rule (~k) i st);
+ | rot_tac (k::ks) i st = rot_tac ks (i+1) (Thm.rotate_rule (~k) i st);
in
fun rot_subgoals_tac (rot, rl) =
rot_tac (if rot then map nNewHyps rl else [])
--- a/src/Provers/classical.ML Sun Jul 05 21:10:23 2009 +0200
+++ b/src/Provers/classical.ML Thu Jul 09 08:55:42 2009 +0200
@@ -209,7 +209,7 @@
fun dup_elim th =
rule_by_tactic (TRYALL (etac revcut_rl))
- ((th RSN (2, revcut_rl)) |> assumption 2 |> Seq.hd);
+ ((th RSN (2, revcut_rl)) |> Thm.assumption 2 |> Seq.hd);
(**** Classical rule sets ****)
--- a/src/Pure/Isar/class.ML Sun Jul 05 21:10:23 2009 +0200
+++ b/src/Pure/Isar/class.ML Thu Jul 09 08:55:42 2009 +0200
@@ -83,14 +83,14 @@
(fst (Locale.intros_of thy class));
(* of_class *)
- val of_class_prop_concl = Logic.mk_inclass (aT, class);
+ val of_class_prop_concl = Logic.mk_of_class (aT, class);
val of_class_prop = case prop of NONE => of_class_prop_concl
| SOME prop => Logic.mk_implies (Morphism.term const_morph
((map_types o map_atyps) (K aT) prop), of_class_prop_concl);
val sup_of_classes = map (snd o rules thy) sups;
val loc_axiom_intros = map Drule.standard' (Locale.axioms_of thy class);
val axclass_intro = #intro (AxClass.get_info thy class);
- val base_sort_trivs = Thm.sort_triv thy (aT, base_sort);
+ val base_sort_trivs = Thm.of_sort (Thm.ctyp_of thy aT, base_sort);
val tac = REPEAT (SOMEGOAL
(Tactic.match_tac (axclass_intro :: sup_of_classes
@ loc_axiom_intros @ base_sort_trivs)
--- a/src/Pure/Isar/code.ML Sun Jul 05 21:10:23 2009 +0200
+++ b/src/Pure/Isar/code.ML Thu Jul 09 08:55:42 2009 +0200
@@ -1,82 +1,69 @@
(* Title: Pure/Isar/code.ML
Author: Florian Haftmann, TU Muenchen
-Abstract executable content of theory. Management of data dependent on
-executable content. Cache assumes non-concurrent processing of a single theory.
+Abstract executable code of theory. Management of data dependent on
+executable code. Cache assumes non-concurrent processing of a single theory.
*)
signature CODE =
sig
+ (*constants*)
+ val check_const: theory -> term -> string
+ val read_bare_const: theory -> string -> string * typ
+ val read_const: theory -> string -> string
+ val string_of_const: theory -> string -> string
+ val args_number: theory -> string -> int
+ val typscheme: theory -> string * typ -> (string * sort) list * typ
+
(*constructor sets*)
val constrset_of_consts: theory -> (string * typ) list
-> string * ((string * sort) list * (string * typ list) list)
- (*typ instantiations*)
- val typscheme: theory -> string * typ -> (string * sort) list * typ
- val inst_thm: theory -> sort Vartab.table -> thm -> thm
-
- (*constants*)
- val string_of_typ: theory -> typ -> string
- val string_of_const: theory -> string -> string
- val no_args: theory -> string -> int
- val check_const: theory -> term -> string
- val read_bare_const: theory -> string -> string * typ
- val read_const: theory -> string -> string
-
(*constant aliasses*)
val add_const_alias: thm -> theory -> theory
val triv_classes: theory -> class list
val resubst_alias: theory -> string -> string
(*code equations*)
- val mk_eqn: theory -> (string -> bool) -> thm * bool -> thm * bool
- val mk_eqn_warning: theory -> (string -> bool) -> thm -> (thm * bool) option
- val mk_eqn_liberal: theory -> (string -> bool) -> thm -> (thm * bool) option
+ val mk_eqn: theory -> thm * bool -> thm * bool
+ val mk_eqn_warning: theory -> thm -> (thm * bool) option
+ val mk_eqn_liberal: theory -> thm -> (thm * bool) option
val assert_eqn: theory -> thm * bool -> thm * bool
val assert_eqns_const: theory -> string
-> (thm * bool) list -> (thm * bool) list
- val const_typ_eqn: thm -> string * typ
- val const_eqn: theory -> thm -> string
+ val const_typ_eqn: theory -> thm -> string * typ
val typscheme_eqn: theory -> thm -> (string * sort) list * typ
val expand_eta: theory -> int -> thm -> thm
- val rewrite_eqn: simpset -> thm -> thm
- val rewrite_head: thm list -> thm -> thm
val norm_args: theory -> thm list -> thm list
val norm_varnames: theory -> thm list -> thm list
- (*case certificates*)
- val case_cert: thm -> string * (int * string list)
-
- (*infrastructure*)
- val add_attribute: string * attribute parser -> theory -> theory
- val purge_data: theory -> theory
-
- (*executable content*)
+ (*executable code*)
val add_datatype: (string * typ) list -> theory -> theory
val add_datatype_cmd: string list -> theory -> theory
val type_interpretation:
(string * ((string * sort) list * (string * typ list) list)
-> theory -> theory) -> theory -> theory
val add_eqn: thm -> theory -> theory
+ val add_eqnl: string * (thm * bool) list lazy -> theory -> theory
val add_nbe_eqn: thm -> theory -> theory
val add_default_eqn: thm -> theory -> theory
val add_default_eqn_attribute: attribute
val add_default_eqn_attrib: Attrib.src
val del_eqn: thm -> theory -> theory
val del_eqns: string -> theory -> theory
- val add_eqnl: string * (thm * bool) list lazy -> theory -> theory
val add_case: thm -> theory -> theory
val add_undefined: string -> theory -> theory
-
- (*data retrieval*)
val get_datatype: theory -> string -> ((string * sort) list * (string * typ list) list)
val get_datatype_of_constr: theory -> string -> string option
- val default_typscheme: theory -> string -> (string * sort) list * typ
val these_eqns: theory -> string -> (thm * bool) list
val all_eqns: theory -> (thm * bool) list
val get_case_scheme: theory -> string -> (int * (int * string list)) option
val undefineds: theory -> string list
val print_codesetup: theory -> unit
+
+ (*infrastructure*)
+ val add_attribute: string * attribute parser -> theory -> theory
+ val purge_data: theory -> theory
end;
signature CODE_DATA_ARGS =
@@ -110,36 +97,36 @@
structure Code : PRIVATE_CODE =
struct
-(* auxiliary *)
+(** auxiliary **)
+
+(* printing *)
fun string_of_typ thy = setmp show_sorts true (Syntax.string_of_typ_global thy);
+
fun string_of_const thy c = case AxClass.inst_of_param thy c
of SOME (c, tyco) => Sign.extern_const thy c ^ " " ^ enclose "[" "]" (Sign.extern_type thy tyco)
| NONE => Sign.extern_const thy c;
-fun no_args thy = length o fst o strip_type o Sign.the_const_type thy;
+
+(* constants *)
+fun check_bare_const thy t = case try dest_Const t
+ of SOME c_ty => c_ty
+ | NONE => error ("Not a constant: " ^ Syntax.string_of_term_global thy t);
-(* utilities *)
+fun check_const thy = AxClass.unoverload_const thy o check_bare_const thy;
+
+fun read_bare_const thy = check_bare_const thy o Syntax.read_term_global thy;
+
+fun read_const thy = AxClass.unoverload_const thy o read_bare_const thy;
fun typscheme thy (c, ty) =
let
val ty' = Logic.unvarifyT ty;
- fun dest (TFree (v, sort)) = (v, sort)
- | dest ty = error ("Illegal type parameter in type scheme: " ^ Syntax.string_of_typ_global thy ty);
- val vs = map dest (Sign.const_typargs thy (c, ty'));
- in (vs, Type.strip_sorts ty') end;
+ in (map dest_TFree (Sign.const_typargs thy (c, ty')), Type.strip_sorts ty') end;
-fun inst_thm thy tvars' thm =
- let
- val tvars = (Term.add_tvars o Thm.prop_of) thm [];
- val inter_sort = Sorts.inter_sort (Sign.classes_of thy);
- fun mk_inst (tvar as (v, sort)) = case Vartab.lookup tvars' v
- of SOME sort' => SOME (pairself (Thm.ctyp_of thy o TVar)
- (tvar, (v, inter_sort (sort, sort'))))
- | NONE => NONE;
- val insts = map_filter mk_inst tvars;
- in Thm.instantiate (insts, []) thm end;
+
+(* code equation transformations *)
fun expand_eta thy k thm =
let
@@ -173,23 +160,6 @@
|> Conv.fconv_rule Drule.beta_eta_conversion
end;
-fun eqn_conv conv =
- let
- fun lhs_conv ct = if can Thm.dest_comb ct
- then (Conv.combination_conv lhs_conv conv) ct
- else Conv.all_conv ct;
- in Conv.combination_conv (Conv.arg_conv lhs_conv) conv end;
-
-fun head_conv conv =
- let
- fun lhs_conv ct = if can Thm.dest_comb ct
- then (Conv.fun_conv lhs_conv) ct
- else conv ct;
- in Conv.fun_conv (Conv.arg_conv lhs_conv) end;
-
-val rewrite_eqn = Conv.fconv_rule o eqn_conv o Simplifier.rewrite;
-val rewrite_head = Conv.fconv_rule o head_conv o MetaSimplifier.rewrite false;
-
fun norm_args thy thms =
let
val num_args_of = length o snd o strip_comb o fst o Logic.dest_equals;
@@ -265,36 +235,272 @@
end;
-(* const aliasses *)
+(** code attributes **)
-structure ConstAlias = TheoryDataFun
-(
- type T = ((string * string) * thm) list * class list;
- val empty = ([], []);
+structure Code_Attr = TheoryDataFun (
+ type T = (string * attribute parser) list;
+ val empty = [];
val copy = I;
val extend = I;
- fun merge _ ((alias1, classes1), (alias2, classes2)) : T =
- (Library.merge (eq_snd Thm.eq_thm_prop) (alias1, alias2),
- Library.merge (op =) (classes1, classes2));
+ fun merge _ = AList.merge (op = : string * string -> bool) (K true);
);
-fun add_const_alias thm thy =
+fun add_attribute (attr as (name, _)) =
+ let
+ fun add_parser ("", parser) attrs = attrs |> rev |> AList.update (op =) ("", parser) |> rev
+ | add_parser (name, parser) attrs = (name, Args.$$$ name |-- parser) :: attrs;
+ in Code_Attr.map (fn attrs => if not (name = "") andalso AList.defined (op =) attrs name
+ then error ("Code attribute " ^ name ^ " already declared") else add_parser attr attrs)
+ end;
+
+val _ = Context.>> (Context.map_theory
+ (Attrib.setup (Binding.name "code")
+ (Scan.peek (fn context =>
+ List.foldr op || Scan.fail (map snd (Code_Attr.get (Context.theory_of context)))))
+ "declare theorems for code generation"));
+
+
+(** data store **)
+
+(* code equations *)
+
+type eqns = bool * (thm * bool) list lazy;
+ (*default flag, theorems with proper flag (perhaps lazy)*)
+
+fun pretty_lthms ctxt r = case Lazy.peek r
+ of SOME thms => map (ProofContext.pretty_thm ctxt o fst) (Exn.release thms)
+ | NONE => [Pretty.str "[...]"];
+
+fun certificate thy f r =
+ case Lazy.peek r
+ of SOME thms => (Lazy.value o f thy) (Exn.release thms)
+ | NONE => let
+ val thy_ref = Theory.check_thy thy;
+ in Lazy.lazy (fn () => (f (Theory.deref thy_ref) o Lazy.force) r) end;
+
+fun add_drop_redundant thy (thm, proper) thms =
+ let
+ val args_of = snd o strip_comb o map_types Type.strip_sorts
+ o fst o Logic.dest_equals o Thm.plain_prop_of;
+ val args = args_of thm;
+ val incr_idx = Logic.incr_indexes ([], Thm.maxidx_of thm + 1);
+ fun matches_args args' = length args <= length args' andalso
+ Pattern.matchess thy (args, (map incr_idx o curry Library.take (length args)) args');
+ fun drop (thm', proper') = if (proper orelse not proper')
+ andalso matches_args (args_of thm') then
+ (warning ("Code generator: dropping redundant code equation\n" ^ Display.string_of_thm thm'); true)
+ else false;
+ in (thm, proper) :: filter_out drop thms end;
+
+fun add_thm thy _ thm (false, thms) = (false, Lazy.map_force (add_drop_redundant thy thm) thms)
+ | add_thm thy true thm (true, thms) = (true, Lazy.map_force (fn thms => thms @ [thm]) thms)
+ | add_thm thy false thm (true, thms) = (false, Lazy.value [thm]);
+
+fun add_lthms lthms _ = (false, lthms);
+
+fun del_thm thm = (apsnd o Lazy.map_force) (remove (eq_fst Thm.eq_thm_prop) (thm, true));
+
+
+(* executable code data *)
+
+datatype spec = Spec of {
+ history_concluded: bool,
+ aliasses: ((string * string) * thm) list * class list,
+ eqns: ((bool * eqns) * (serial * eqns) list) Symtab.table
+ (*with explicit history*),
+ dtyps: ((serial * ((string * sort) list * (string * typ list) list)) list) Symtab.table
+ (*with explicit history*),
+ cases: (int * (int * string list)) Symtab.table * unit Symtab.table
+};
+
+fun make_spec ((history_concluded, aliasses), (eqns, (dtyps, cases))) =
+ Spec { history_concluded = history_concluded, aliasses = aliasses,
+ eqns = eqns, dtyps = dtyps, cases = cases };
+fun map_spec f (Spec { history_concluded = history_concluded, aliasses = aliasses, eqns = eqns,
+ dtyps = dtyps, cases = cases }) =
+ make_spec (f ((history_concluded, aliasses), (eqns, (dtyps, cases))));
+fun merge_spec (Spec { history_concluded = _, aliasses = aliasses1, eqns = eqns1,
+ dtyps = dtyps1, cases = (cases1, undefs1) },
+ Spec { history_concluded = _, aliasses = aliasses2, eqns = eqns2,
+ dtyps = dtyps2, cases = (cases2, undefs2) }) =
let
- val lhs_rhs = case try Logic.dest_equals (Thm.prop_of thm)
- of SOME lhs_rhs => lhs_rhs
- | _ => error ("Not an equation: " ^ Display.string_of_thm thm);
- val c_c' = case try (pairself (AxClass.unoverload_const thy o dest_Const)) lhs_rhs
- of SOME c_c' => c_c'
- | _ => error ("Not an equation with two constants: " ^ Display.string_of_thm thm);
- val some_class = the_list (AxClass.class_of_param thy (snd c_c'));
- in thy |>
- ConstAlias.map (fn (alias, classes) =>
- ((c_c', thm) :: alias, fold (insert (op =)) some_class classes))
- end;
+ val aliasses = (Library.merge (eq_snd Thm.eq_thm_prop) (pairself fst (aliasses1, aliasses2)),
+ Library.merge (op =) (pairself snd (aliasses1, aliasses2)));
+ fun merge_eqns ((_, history1), (_, history2)) =
+ let
+ val raw_history = AList.merge (op = : serial * serial -> bool)
+ (K true) (history1, history2)
+ val filtered_history = filter_out (fst o snd) raw_history
+ val history = if null filtered_history
+ then raw_history else filtered_history;
+ in ((false, (snd o hd) history), history) end;
+ val eqns = Symtab.join (K merge_eqns) (eqns1, eqns2);
+ val dtyps = Symtab.join (K (AList.merge (op =) (K true))) (dtyps1, dtyps2);
+ val cases = (Symtab.merge (K true) (cases1, cases2),
+ Symtab.merge (K true) (undefs1, undefs2));
+ in make_spec ((false, aliasses), (eqns, (dtyps, cases))) end;
+
+fun history_concluded (Spec { history_concluded, ... }) = history_concluded;
+fun the_aliasses (Spec { aliasses, ... }) = aliasses;
+fun the_eqns (Spec { eqns, ... }) = eqns;
+fun the_dtyps (Spec { dtyps, ... }) = dtyps;
+fun the_cases (Spec { cases, ... }) = cases;
+val map_history_concluded = map_spec o apfst o apfst;
+val map_aliasses = map_spec o apfst o apsnd;
+val map_eqns = map_spec o apsnd o apfst;
+val map_dtyps = map_spec o apsnd o apsnd o apfst;
+val map_cases = map_spec o apsnd o apsnd o apsnd;
+
+
+(* data slots dependent on executable code *)
+
+(*private copy avoids potential conflict of table exceptions*)
+structure Datatab = TableFun(type key = int val ord = int_ord);
+
+local
+
+type kind = {
+ empty: Object.T,
+ purge: theory -> string list -> Object.T -> Object.T
+};
+
+val kinds = ref (Datatab.empty: kind Datatab.table);
+val kind_keys = ref ([]: serial list);
+
+fun invoke f k = case Datatab.lookup (! kinds) k
+ of SOME kind => f kind
+ | NONE => sys_error "Invalid code data identifier";
+
+in
+
+fun declare_data empty purge =
+ let
+ val k = serial ();
+ val kind = {empty = empty, purge = purge};
+ val _ = change kinds (Datatab.update (k, kind));
+ val _ = change kind_keys (cons k);
+ in k end;
+
+fun invoke_init k = invoke (fn kind => #empty kind) k;
+
+fun invoke_purge_all thy cs =
+ fold (fn k => Datatab.map_entry k
+ (invoke (fn kind => #purge kind thy cs) k)) (! kind_keys);
+
+end; (*local*)
+
+
+(* theory store *)
+
+local
+
+type data = Object.T Datatab.table;
+val empty_data = Datatab.empty : data;
+
+structure Code_Data = TheoryDataFun
+(
+ type T = spec * data ref;
+ val empty = (make_spec ((false, ([], [])),
+ (Symtab.empty, (Symtab.empty, (Symtab.empty, Symtab.empty)))), ref empty_data);
+ fun copy (spec, data) = (spec, ref (! data));
+ val extend = copy;
+ fun merge pp ((spec1, data1), (spec2, data2)) =
+ (merge_spec (spec1, spec2), ref empty_data);
+);
+
+fun thy_data f thy = f ((snd o Code_Data.get) thy);
+
+fun get_ensure_init kind data_ref =
+ case Datatab.lookup (! data_ref) kind
+ of SOME x => x
+ | NONE => let val y = invoke_init kind
+ in (change data_ref (Datatab.update (kind, y)); y) end;
+
+in
+
+(* access to executable code *)
+
+val the_exec = fst o Code_Data.get;
+
+fun complete_class_params thy cs =
+ fold (fn c => case AxClass.inst_of_param thy c
+ of NONE => insert (op =) c
+ | SOME (c', _) => insert (op =) c' #> insert (op =) c) cs [];
+
+fun map_exec_purge touched f thy =
+ Code_Data.map (fn (exec, data) => (f exec, ref (case touched
+ of SOME cs => invoke_purge_all thy (complete_class_params thy cs) (! data)
+ | NONE => empty_data))) thy;
+
+val purge_data = (Code_Data.map o apsnd) (K (ref empty_data));
+
+fun change_eqns delete c f = (map_exec_purge (SOME [c]) o map_eqns
+ o (if delete then Symtab.map_entry c else Symtab.map_default (c, ((false, (true, Lazy.value [])), [])))
+ o apfst) (fn (_, eqns) => (true, f eqns));
+
+fun del_eqns c = change_eqns true c (K (false, Lazy.value []));
+
+
+(* tackling equation history *)
+
+fun get_eqns thy c =
+ Symtab.lookup ((the_eqns o the_exec) thy) c
+ |> Option.map (Lazy.force o snd o snd o fst)
+ |> these;
+
+fun continue_history thy = if (history_concluded o the_exec) thy
+ then thy
+ |> (Code_Data.map o apfst o map_history_concluded) (K false)
+ |> SOME
+ else NONE;
+
+fun conclude_history thy = if (history_concluded o the_exec) thy
+ then NONE
+ else thy
+ |> (Code_Data.map o apfst)
+ ((map_eqns o Symtab.map) (fn ((changed, current), history) =>
+ ((false, current),
+ if changed then (serial (), current) :: history else history))
+ #> map_history_concluded (K true))
+ |> SOME;
+
+val _ = Context.>> (Context.map_theory (Code_Data.init
+ #> Theory.at_begin continue_history
+ #> Theory.at_end conclude_history));
+
+
+(* access to data dependent on abstract executable code *)
+
+fun get_data (kind, _, dest) = thy_data (get_ensure_init kind #> dest);
+
+fun change_data (kind, mk, dest) =
+ let
+ fun chnge data_ref f =
+ let
+ val data = get_ensure_init kind data_ref;
+ val data' = f (dest data);
+ in (change data_ref (Datatab.update (kind, mk data')); data') end;
+ in thy_data chnge end;
+
+fun change_yield_data (kind, mk, dest) =
+ let
+ fun chnge data_ref f =
+ let
+ val data = get_ensure_init kind data_ref;
+ val (x, data') = f (dest data);
+ in (x, (change data_ref (Datatab.update (kind, mk data')); data')) end;
+ in thy_data chnge end;
+
+end; (*local*)
+
+
+(** retrieval interfaces **)
+
+(* constant aliasses *)
fun resubst_alias thy =
let
- val alias = fst (ConstAlias.get thy);
+ val alias = (fst o the_aliasses o the_exec) thy;
val subst_inst_param = Option.map fst o AxClass.inst_of_param thy;
fun subst_alias c =
get_first (fn ((c', c''), _) => if c = c'' then SOME c' else NONE) alias;
@@ -303,23 +509,17 @@
#> perhaps subst_alias
end;
-val triv_classes = snd o ConstAlias.get;
+val triv_classes = snd o the_aliasses o the_exec;
-(* reading constants as terms *)
-
-fun check_bare_const thy t = case try dest_Const t
- of SOME c_ty => c_ty
- | NONE => error ("Not a constant: " ^ Syntax.string_of_term_global thy t);
+(** foundation **)
-fun check_const thy = AxClass.unoverload_const thy o check_bare_const thy;
+(* constants *)
-fun read_bare_const thy = check_bare_const thy o Syntax.read_term_global thy;
-
-fun read_const thy = AxClass.unoverload_const thy o read_bare_const thy;
+fun args_number thy = length o fst o strip_type o Sign.the_const_type thy;
-(* constructor sets *)
+(* datatypes *)
fun constrset_of_consts thy cs =
let
@@ -359,6 +559,22 @@
val cs''' = map (inst vs) cs'';
in (tyco, (vs, rev cs''')) end;
+fun get_datatype thy tyco =
+ case these (Symtab.lookup ((the_dtyps o the_exec) thy) tyco)
+ of (_, spec) :: _ => spec
+ | [] => Sign.arity_number thy tyco
+ |> Name.invents Name.context Name.aT
+ |> map (rpair [])
+ |> rpair [];
+
+fun get_datatype_of_constr thy c =
+ case (snd o strip_type o Sign.the_const_type thy) c
+ of Type (tyco, _) => if member (op =) ((map fst o snd o get_datatype thy) tyco) c
+ then SOME tyco else NONE
+ | _ => NONE;
+
+fun is_constr thy = is_some o get_datatype_of_constr thy;
+
(* code equations *)
@@ -373,7 +589,7 @@
in not (has_duplicates (op =) ((fold o fold_aterms)
(fn Var (v, _) => cons v | _ => I) args [])) end;
-fun gen_assert_eqn thy is_constr_head is_constr_pat (thm, proper) =
+fun gen_assert_eqn thy is_constr_pat (thm, proper) =
let
val (lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm
handle TERM _ => bad_thm ("Not an equation: " ^ Display.string_of_thm thm)
@@ -425,7 +641,7 @@
then ()
else bad_thm ("Polymorphic constant as head in equation\n"
^ Display.string_of_thm thm)
- val _ = if not (is_constr_head c)
+ val _ = if not (is_constr thy c)
then ()
else bad_thm ("Constructor as head in equation\n"
^ Display.string_of_thm thm)
@@ -438,35 +654,70 @@
^ string_of_typ thy ty_decl)
in (thm, proper) end;
-fun assert_eqn thy is_constr = error_thm (gen_assert_eqn thy is_constr is_constr);
+fun assert_eqn thy = error_thm (gen_assert_eqn thy (is_constr thy));
+
+fun meta_rewrite thy = LocalDefs.meta_rewrite_rule (ProofContext.init thy);
-val const_typ_eqn = dest_Const o fst o strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of;
+fun mk_eqn thy = error_thm (gen_assert_eqn thy (K true)) o
+ apfst (meta_rewrite thy);
+fun mk_eqn_warning thy = Option.map (fn (thm, _) => (thm, is_linear thm))
+ o warning_thm (gen_assert_eqn thy (K true)) o rpair false o meta_rewrite thy;
+
+fun mk_eqn_liberal thy = Option.map (fn (thm, _) => (thm, is_linear thm))
+ o try_thm (gen_assert_eqn thy (K true)) o rpair false o meta_rewrite thy;
(*those following are permissive wrt. to overloaded constants!*)
-fun mk_eqn thy is_constr_head = error_thm (gen_assert_eqn thy is_constr_head (K true)) o
- apfst (LocalDefs.meta_rewrite_rule (ProofContext.init thy));
-
-fun mk_eqn_warning thy is_constr_head = Option.map (fn (thm, _) => (thm, is_linear thm))
- o warning_thm (gen_assert_eqn thy is_constr_head (K true))
- o rpair false o LocalDefs.meta_rewrite_rule (ProofContext.init thy);
-
-fun mk_eqn_liberal thy is_constr_head = Option.map (fn (thm, _) => (thm, is_linear thm))
- o try_thm (gen_assert_eqn thy is_constr_head (K true))
- o rpair false o LocalDefs.meta_rewrite_rule (ProofContext.init thy);
-
-fun const_typ_eqn_unoverload thy thm =
+fun const_typ_eqn thy thm =
let
- val (c, ty) = const_typ_eqn thm;
+ val (c, ty) = (dest_Const o fst o strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of) thm;
val c' = AxClass.unoverload_const thy (c, ty);
in (c', ty) end;
-fun typscheme_eqn thy = typscheme thy o const_typ_eqn_unoverload thy;
-fun const_eqn thy = fst o const_typ_eqn_unoverload thy;
+fun typscheme_eqn thy = typscheme thy o const_typ_eqn thy;
+fun const_eqn thy = fst o const_typ_eqn thy;
+
+fun assert_eqns_const thy c eqns =
+ let
+ fun cert (eqn as (thm, _)) = if c = const_eqn thy thm
+ then eqn else error ("Wrong head of code equation,\nexpected constant "
+ ^ string_of_const thy c ^ "\n" ^ Display.string_of_thm thm)
+ in map (cert o assert_eqn thy) eqns end;
+
+fun common_typ_eqns thy [] = []
+ | common_typ_eqns thy [thm] = [thm]
+ | common_typ_eqns thy (thms as thm :: _) = (*FIXME is too general*)
+ let
+ fun incr_thm thm max =
+ let
+ val thm' = incr_indexes max thm;
+ val max' = Thm.maxidx_of thm' + 1;
+ in (thm', max') end;
+ val (thms', maxidx) = fold_map incr_thm thms 0;
+ val ty1 :: tys = map (snd o const_typ_eqn thy) thms';
+ fun unify ty env = Sign.typ_unify thy (ty1, ty) env
+ handle Type.TUNIFY =>
+ error ("Type unificaton failed, while unifying code equations\n"
+ ^ (cat_lines o map Display.string_of_thm) thms
+ ^ "\nwith types\n"
+ ^ (cat_lines o map (string_of_typ thy)) (ty1 :: tys));
+ val (env, _) = fold unify tys (Vartab.empty, maxidx)
+ val instT = Vartab.fold (fn (x_i, (sort, ty)) =>
+ cons (Thm.ctyp_of thy (TVar (x_i, sort)), Thm.ctyp_of thy ty)) env [];
+ in map (Thm.instantiate (instT, [])) thms' end;
+
+fun these_eqns thy c =
+ get_eqns thy c
+ |> (map o apfst) (Thm.transfer thy)
+ |> burrow_fst (common_typ_eqns thy);
+
+fun all_eqns thy =
+ Symtab.dest ((the_eqns o the_exec) thy)
+ |> maps (Lazy.force o snd o snd o fst o snd);
-(* case cerificates *)
+(* cases *)
fun case_certificate thm =
let
@@ -508,255 +759,12 @@
handle Bind => error "bad case certificate"
| TERM _ => error "bad case certificate";
-
-(** code attributes **)
-
-structure CodeAttr = TheoryDataFun (
- type T = (string * attribute parser) list;
- val empty = [];
- val copy = I;
- val extend = I;
- fun merge _ = AList.merge (op = : string * string -> bool) (K true);
-);
-
-fun add_attribute (attr as (name, _)) =
- let
- fun add_parser ("", parser) attrs = attrs |> rev |> AList.update (op =) ("", parser) |> rev
- | add_parser (name, parser) attrs = (name, Args.$$$ name |-- parser) :: attrs;
- in CodeAttr.map (fn attrs => if not (name = "") andalso AList.defined (op =) attrs name
- then error ("Code attribute " ^ name ^ " already declared") else add_parser attr attrs)
- end;
-
-val _ = Context.>> (Context.map_theory
- (Attrib.setup (Binding.name "code")
- (Scan.peek (fn context =>
- List.foldr op || Scan.fail (map snd (CodeAttr.get (Context.theory_of context)))))
- "declare theorems for code generation"));
-
-
-
-(** logical and syntactical specification of executable code **)
-
-(* code equations *)
-
-type eqns = bool * (thm * bool) list lazy;
- (*default flag, theorems with proper flag (perhaps lazy)*)
-
-fun pretty_lthms ctxt r = case Lazy.peek r
- of SOME thms => map (ProofContext.pretty_thm ctxt o fst) (Exn.release thms)
- | NONE => [Pretty.str "[...]"];
-
-fun certificate thy f r =
- case Lazy.peek r
- of SOME thms => (Lazy.value o f thy) (Exn.release thms)
- | NONE => let
- val thy_ref = Theory.check_thy thy;
- in Lazy.lazy (fn () => (f (Theory.deref thy_ref) o Lazy.force) r) end;
+fun get_case_scheme thy = Symtab.lookup ((fst o the_cases o the_exec) thy);
-fun add_drop_redundant thy (thm, proper) thms =
- let
- val args_of = snd o strip_comb o map_types Type.strip_sorts
- o fst o Logic.dest_equals o Thm.plain_prop_of;
- val args = args_of thm;
- val incr_idx = Logic.incr_indexes ([], Thm.maxidx_of thm + 1);
- fun matches_args args' = length args <= length args' andalso
- Pattern.matchess thy (args, (map incr_idx o curry Library.take (length args)) args');
- fun drop (thm', proper') = if (proper orelse not proper')
- andalso matches_args (args_of thm') then
- (warning ("Code generator: dropping redundant code equation\n" ^ Display.string_of_thm thm'); true)
- else false;
- in (thm, proper) :: filter_out drop thms end;
-
-fun add_thm thy _ thm (false, thms) = (false, Lazy.map_force (add_drop_redundant thy thm) thms)
- | add_thm thy true thm (true, thms) = (true, Lazy.map_force (fn thms => thms @ [thm]) thms)
- | add_thm thy false thm (true, thms) = (false, Lazy.value [thm]);
-
-fun add_lthms lthms _ = (false, lthms);
-
-fun del_thm thm = (apsnd o Lazy.map_force) (remove (eq_fst Thm.eq_thm_prop) (thm, true));
-
-
-(* specification data *)
-
-datatype spec = Spec of {
- concluded_history: bool,
- eqns: ((bool * eqns) * (serial * eqns) list) Symtab.table
- (*with explicit history*),
- dtyps: ((serial * ((string * sort) list * (string * typ list) list)) list) Symtab.table
- (*with explicit history*),
- cases: (int * (int * string list)) Symtab.table * unit Symtab.table
-};
-
-fun make_spec ((concluded_history, eqns), (dtyps, cases)) =
- Spec { concluded_history = concluded_history, eqns = eqns, dtyps = dtyps, cases = cases };
-fun map_spec f (Spec { concluded_history = concluded_history, eqns = eqns,
- dtyps = dtyps, cases = cases }) =
- make_spec (f ((concluded_history, eqns), (dtyps, cases)));
-fun merge_spec (Spec { concluded_history = _, eqns = eqns1, dtyps = dtyps1, cases = (cases1, undefs1) },
- Spec { concluded_history = _, eqns = eqns2, dtyps = dtyps2, cases = (cases2, undefs2) }) =
- let
- fun merge_eqns ((_, history1), (_, history2)) =
- let
- val raw_history = AList.merge (op = : serial * serial -> bool)
- (K true) (history1, history2)
- val filtered_history = filter_out (fst o snd) raw_history
- val history = if null filtered_history
- then raw_history else filtered_history;
- in ((false, (snd o hd) history), history) end;
- val eqns = Symtab.join (K merge_eqns) (eqns1, eqns2);
- val dtyps = Symtab.join (K (AList.merge (op =) (K true))) (dtyps1, dtyps2);
- val cases = (Symtab.merge (K true) (cases1, cases2),
- Symtab.merge (K true) (undefs1, undefs2));
- in make_spec ((false, eqns), (dtyps, cases)) end;
-
-
-(* code setup data *)
-
-fun the_spec (Spec x) = x;
-fun the_eqns (Spec { eqns, ... }) = eqns;
-fun the_dtyps (Spec { dtyps, ... }) = dtyps;
-fun the_cases (Spec { cases, ... }) = cases;
-fun history_concluded (Spec { concluded_history, ... }) = concluded_history;
-val map_concluded_history = map_spec o apfst o apfst;
-val map_eqns = map_spec o apfst o apsnd;
-val map_dtyps = map_spec o apsnd o apfst;
-val map_cases = map_spec o apsnd o apsnd;
+val undefineds = Symtab.keys o snd o the_cases o the_exec;
-(* data slots dependent on executable content *)
-
-(*private copy avoids potential conflict of table exceptions*)
-structure Datatab = TableFun(type key = int val ord = int_ord);
-
-local
-
-type kind = {
- empty: Object.T,
- purge: theory -> string list -> Object.T -> Object.T
-};
-
-val kinds = ref (Datatab.empty: kind Datatab.table);
-val kind_keys = ref ([]: serial list);
-
-fun invoke f k = case Datatab.lookup (! kinds) k
- of SOME kind => f kind
- | NONE => sys_error "Invalid code data identifier";
-
-in
-
-fun declare_data empty purge =
- let
- val k = serial ();
- val kind = {empty = empty, purge = purge};
- val _ = change kinds (Datatab.update (k, kind));
- val _ = change kind_keys (cons k);
- in k end;
-
-fun invoke_init k = invoke (fn kind => #empty kind) k;
-
-fun invoke_purge_all thy cs =
- fold (fn k => Datatab.map_entry k
- (invoke (fn kind => #purge kind thy cs) k)) (! kind_keys);
-
-end; (*local*)
-
-
-(** theory store **)
-
-local
-
-type data = Object.T Datatab.table;
-val empty_data = Datatab.empty : data;
-
-structure Code_Data = TheoryDataFun
-(
- type T = spec * data ref;
- val empty = (make_spec ((false, Symtab.empty),
- (Symtab.empty, (Symtab.empty, Symtab.empty))), ref empty_data);
- fun copy (spec, data) = (spec, ref (! data));
- val extend = copy;
- fun merge pp ((spec1, data1), (spec2, data2)) =
- (merge_spec (spec1, spec2), ref empty_data);
-);
-
-fun thy_data f thy = f ((snd o Code_Data.get) thy);
-
-fun get_ensure_init kind data_ref =
- case Datatab.lookup (! data_ref) kind
- of SOME x => x
- | NONE => let val y = invoke_init kind
- in (change data_ref (Datatab.update (kind, y)); y) end;
-
-in
-
-(* access to executable content *)
-
-val the_exec = fst o Code_Data.get;
-
-fun complete_class_params thy cs =
- fold (fn c => case AxClass.inst_of_param thy c
- of NONE => insert (op =) c
- | SOME (c', _) => insert (op =) c' #> insert (op =) c) cs [];
-
-fun map_exec_purge touched f thy =
- Code_Data.map (fn (exec, data) => (f exec, ref (case touched
- of SOME cs => invoke_purge_all thy (complete_class_params thy cs) (! data)
- | NONE => empty_data))) thy;
-
-val purge_data = (Code_Data.map o apsnd) (K (ref empty_data));
-
-
-(* tackling equation history *)
-
-fun get_eqns thy c =
- Symtab.lookup ((the_eqns o the_exec) thy) c
- |> Option.map (Lazy.force o snd o snd o fst)
- |> these;
-
-fun continue_history thy = if (history_concluded o the_exec) thy
- then thy
- |> (Code_Data.map o apfst o map_concluded_history) (K false)
- |> SOME
- else NONE;
-
-fun conclude_history thy = if (history_concluded o the_exec) thy
- then NONE
- else thy
- |> (Code_Data.map o apfst)
- ((map_eqns o Symtab.map) (fn ((changed, current), history) =>
- ((false, current),
- if changed then (serial (), current) :: history else history))
- #> map_concluded_history (K true))
- |> SOME;
-
-val _ = Context.>> (Context.map_theory (Code_Data.init
- #> Theory.at_begin continue_history
- #> Theory.at_end conclude_history));
-
-
-(* access to data dependent on abstract executable content *)
-
-fun get_data (kind, _, dest) = thy_data (get_ensure_init kind #> dest);
-
-fun change_data (kind, mk, dest) =
- let
- fun chnge data_ref f =
- let
- val data = get_ensure_init kind data_ref;
- val data' = f (dest data);
- in (change data_ref (Datatab.update (kind, mk data')); data') end;
- in thy_data chnge end;
-
-fun change_yield_data (kind, mk, dest) =
- let
- fun chnge data_ref f =
- let
- val data = get_ensure_init kind data_ref;
- val (x, data') = f (dest data);
- in (x, (change data_ref (Datatab.update (kind, mk data')); data')) end;
- in thy_data chnge end;
-
-end; (*local*)
+(* diagnostic *)
fun print_codesetup thy =
let
@@ -787,7 +795,7 @@
val dtyps = the_dtyps exec
|> Symtab.dest
|> map (fn (dtco, (_, (vs, cos)) :: _) =>
- (Syntax.string_of_typ_global thy (Type (dtco, map TFree vs)), cos))
+ (string_of_typ thy (Type (dtco, map TFree vs)), cos))
|> sort (string_ord o pairself fst)
in
(Pretty.writeln o Pretty.chunks) [
@@ -805,102 +813,41 @@
end;
-(** theorem transformation and certification **)
+(** declaring executable ingredients **)
+
+(* constant aliasses *)
-fun common_typ_eqns thy [] = []
- | common_typ_eqns thy [thm] = [thm]
- | common_typ_eqns thy (thms as thm :: _) = (*FIXME is too general*)
- let
- fun incr_thm thm max =
- let
- val thm' = incr_indexes max thm;
- val max' = Thm.maxidx_of thm' + 1;
- in (thm', max') end;
- val (thms', maxidx) = fold_map incr_thm thms 0;
- val ty1 :: tys = map (snd o const_typ_eqn) thms';
- fun unify ty env = Sign.typ_unify thy (ty1, ty) env
- handle Type.TUNIFY =>
- error ("Type unificaton failed, while unifying code equations\n"
- ^ (cat_lines o map Display.string_of_thm) thms
- ^ "\nwith types\n"
- ^ (cat_lines o map (string_of_typ thy)) (ty1 :: tys));
- val (env, _) = fold unify tys (Vartab.empty, maxidx)
- val instT = Vartab.fold (fn (x_i, (sort, ty)) =>
- cons (Thm.ctyp_of thy (TVar (x_i, sort)), Thm.ctyp_of thy ty)) env [];
- in map (Thm.instantiate (instT, [])) thms' end;
+fun add_const_alias thm thy =
+ let
+ val (ofclass, eqn) = case try Logic.dest_equals (Thm.prop_of thm)
+ of SOME ofclass_eq => ofclass_eq
+ | _ => error ("Bad certificate: " ^ Display.string_of_thm thm);
+ val (T, class) = case try Logic.dest_of_class ofclass
+ of SOME T_class => T_class
+ | _ => error ("Bad certificate: " ^ Display.string_of_thm thm);
+ val tvar = case try Term.dest_TVar T
+ of SOME tvar => tvar
+ | _ => error ("Bad type: " ^ Display.string_of_thm thm);
+ val _ = if Term.add_tvars eqn [] = [tvar] then ()
+ else error ("Inconsistent type: " ^ Display.string_of_thm thm);
+ val lhs_rhs = case try Logic.dest_equals eqn
+ of SOME lhs_rhs => lhs_rhs
+ | _ => error ("Not an equation: " ^ Syntax.string_of_term_global thy eqn);
+ val c_c' = case try (pairself (check_const thy)) lhs_rhs
+ of SOME c_c' => c_c'
+ | _ => error ("Not an equation with two constants: "
+ ^ Syntax.string_of_term_global thy eqn);
+ val _ = if the_list (AxClass.class_of_param thy (snd c_c')) = [class] then ()
+ else error ("Inconsistent class: " ^ Display.string_of_thm thm);
+ in thy |>
+ (map_exec_purge NONE o map_aliasses) (fn (alias, classes) =>
+ ((c_c', thm) :: alias, insert (op =) class classes))
+ end;
-(** interfaces and attributes **)
-
-fun get_datatype thy tyco =
- case these (Symtab.lookup ((the_dtyps o the_exec) thy) tyco)
- of (_, spec) :: _ => spec
- | [] => Sign.arity_number thy tyco
- |> Name.invents Name.context Name.aT
- |> map (rpair [])
- |> rpair [];
-
-fun get_datatype_of_constr thy c =
- case (snd o strip_type o Sign.the_const_type thy) c
- of Type (tyco, _) => if member (op =) ((map fst o snd o get_datatype thy) tyco) c
- then SOME tyco else NONE
- | _ => NONE;
-
-fun is_constr thy = is_some o get_datatype_of_constr thy;
-
-val assert_eqn = fn thy => assert_eqn thy (is_constr thy);
-
-fun assert_eqns_const thy c eqns =
- let
- fun cert (eqn as (thm, _)) = if c = const_eqn thy thm
- then eqn else error ("Wrong head of code equation,\nexpected constant "
- ^ string_of_const thy c ^ "\n" ^ Display.string_of_thm thm)
- in map (cert o assert_eqn thy) eqns end;
-
-fun change_eqns delete c f = (map_exec_purge (SOME [c]) o map_eqns
- o (if delete then Symtab.map_entry c else Symtab.map_default (c, ((false, (true, Lazy.value [])), [])))
- o apfst) (fn (_, eqns) => (true, f eqns));
-
-fun gen_add_eqn default (eqn as (thm, _)) thy =
- let val c = const_eqn thy thm
- in change_eqns false c (add_thm thy default eqn) thy end;
+(* datatypes *)
-fun add_eqn thm thy =
- gen_add_eqn false (mk_eqn thy (is_constr thy) (thm, true)) thy;
-
-fun add_warning_eqn thm thy =
- case mk_eqn_warning thy (is_constr thy) thm
- of SOME eqn => gen_add_eqn false eqn thy
- | NONE => thy;
-
-fun add_default_eqn thm thy =
- case mk_eqn_liberal thy (is_constr thy) thm
- of SOME eqn => gen_add_eqn true eqn thy
- | NONE => thy;
-
-fun add_nbe_eqn thm thy =
- gen_add_eqn false (mk_eqn thy (is_constr thy) (thm, false)) thy;
-
-fun add_eqnl (c, lthms) thy =
- let
- val lthms' = certificate thy (fn thy => assert_eqns_const thy c) lthms;
- in change_eqns false c (add_lthms lthms') thy end;
-
-val add_default_eqn_attribute = Thm.declaration_attribute
- (fn thm => Context.mapping (add_default_eqn thm) I);
-val add_default_eqn_attrib = Attrib.internal (K add_default_eqn_attribute);
-
-fun del_eqn thm thy = case mk_eqn_liberal thy (is_constr thy) thm
- of SOME (thm, _) => change_eqns true (const_eqn thy thm) (del_thm thm) thy
- | NONE => thy;
-
-fun del_eqns c = change_eqns true c (K (false, Lazy.value []));
-
-fun get_case_scheme thy = Symtab.lookup ((fst o the_cases o the_exec) thy);
-
-val undefineds = Symtab.keys o snd o the_cases o the_exec;
-
-structure TypeInterpretation = InterpretationFun(type T = string * serial val eq = eq_snd (op =) : T * T -> bool);
+structure Type_Interpretation = InterpretationFun(type T = string * serial val eq = eq_snd (op =) : T * T -> bool);
fun add_datatype raw_cs thy =
let
@@ -917,10 +864,10 @@
|> map_exec_purge NONE
((map_dtyps o Symtab.map_default (tyco, [])) (cons (serial (), vs_cos))
#> (map_cases o apfst) drop_outdated_cases)
- |> TypeInterpretation.data (tyco, serial ())
+ |> Type_Interpretation.data (tyco, serial ())
end;
-fun type_interpretation f = TypeInterpretation.interpretation
+fun type_interpretation f = Type_Interpretation.interpretation
(fn (tyco, _) => fn thy => f (tyco, get_datatype thy tyco) thy);
fun add_datatype_cmd raw_cs thy =
@@ -928,6 +875,59 @@
val cs = map (read_bare_const thy) raw_cs;
in add_datatype cs thy end;
+
+(* code equations *)
+
+fun gen_add_eqn default (eqn as (thm, _)) thy =
+ let val c = const_eqn thy thm
+ in change_eqns false c (add_thm thy default eqn) thy end;
+
+fun add_eqn thm thy =
+ gen_add_eqn false (mk_eqn thy (thm, true)) thy;
+
+fun add_warning_eqn thm thy =
+ case mk_eqn_warning thy thm
+ of SOME eqn => gen_add_eqn false eqn thy
+ | NONE => thy;
+
+fun add_default_eqn thm thy =
+ case mk_eqn_liberal thy thm
+ of SOME eqn => gen_add_eqn true eqn thy
+ | NONE => thy;
+
+fun add_nbe_eqn thm thy =
+ gen_add_eqn false (mk_eqn thy (thm, false)) thy;
+
+fun add_eqnl (c, lthms) thy =
+ let
+ val lthms' = certificate thy (fn thy => assert_eqns_const thy c) lthms;
+ in change_eqns false c (add_lthms lthms') thy end;
+
+val add_default_eqn_attribute = Thm.declaration_attribute
+ (fn thm => Context.mapping (add_default_eqn thm) I);
+val add_default_eqn_attrib = Attrib.internal (K add_default_eqn_attribute);
+
+fun del_eqn thm thy = case mk_eqn_liberal thy thm
+ of SOME (thm, _) => change_eqns true (const_eqn thy thm) (del_thm thm) thy
+ | NONE => thy;
+
+val _ = Context.>> (Context.map_theory
+ (let
+ fun mk_attribute f = Thm.declaration_attribute (fn thm => Context.mapping (f thm) I);
+ fun add_simple_attribute (name, f) =
+ add_attribute (name, Scan.succeed (mk_attribute f));
+ fun add_del_attribute (name, (add, del)) =
+ add_attribute (name, Args.del |-- Scan.succeed (mk_attribute del)
+ || Scan.succeed (mk_attribute add))
+ in
+ Type_Interpretation.init
+ #> add_del_attribute ("", (add_warning_eqn, del_eqn))
+ #> add_simple_attribute ("nbe", add_nbe_eqn)
+ end));
+
+
+(* cases *)
+
fun add_case thm thy =
let
val (c, (k, case_pats)) = case_cert thm;
@@ -940,48 +940,12 @@
fun add_undefined c thy =
(map_exec_purge (SOME [c]) o map_cases o apsnd) (Symtab.update (c, ())) thy;
-val _ = Context.>> (Context.map_theory
- (let
- fun mk_attribute f = Thm.declaration_attribute (fn thm => Context.mapping (f thm) I);
- fun add_simple_attribute (name, f) =
- add_attribute (name, Scan.succeed (mk_attribute f));
- fun add_del_attribute (name, (add, del)) =
- add_attribute (name, Args.del |-- Scan.succeed (mk_attribute del)
- || Scan.succeed (mk_attribute add))
- in
- TypeInterpretation.init
- #> add_del_attribute ("", (add_warning_eqn, del_eqn))
- #> add_simple_attribute ("nbe", add_nbe_eqn)
- end));
-
-fun these_eqns thy c =
- get_eqns thy c
- |> (map o apfst) (Thm.transfer thy)
- |> burrow_fst (common_typ_eqns thy);
-
-fun all_eqns thy =
- Symtab.dest ((the_eqns o the_exec) thy)
- |> maps (Lazy.force o snd o snd o fst o snd);
-
-fun default_typscheme thy c =
- let
- fun the_const_typscheme c = (curry (typscheme thy) c o snd o dest_Const
- o TermSubst.zero_var_indexes o curry Const "" o Sign.the_const_type thy) c;
- fun strip_sorts (vs, ty) = (map (fn (v, _) => (v, [])) vs, ty);
- in case AxClass.class_of_param thy c
- of SOME class => ([(Name.aT, [class])], snd (the_const_typscheme c))
- | NONE => if is_constr thy c
- then strip_sorts (the_const_typscheme c)
- else case get_eqns thy c
- of (thm, _) :: _ => (typscheme_eqn thy o Drule.zero_var_indexes) thm
- | [] => strip_sorts (the_const_typscheme c) end;
-
end; (*struct*)
-(** type-safe interfaces for data depedent on executable content **)
+(** type-safe interfaces for data depedent on executable code **)
-functor CodeDataFun(Data: CODE_DATA_ARGS): CODE_DATA =
+functor Code_Data_Fun(Data: CODE_DATA_ARGS): CODE_DATA =
struct
type T = Data.T;
--- a/src/Pure/Isar/rule_insts.ML Sun Jul 05 21:10:23 2009 +0200
+++ b/src/Pure/Isar/rule_insts.ML Thu Jul 09 08:55:42 2009 +0200
@@ -354,7 +354,7 @@
instantiate ([], [(cvar ("V", 0), cvar ("V", maxidx + 1)),
(cvar ("W", 0), cvar ("W", maxidx + 1))]) Drule.revcut_rl;
in
- (case Seq.list_of (bicompose false (false, rl, Thm.nprems_of rl) 1 revcut_rl') of
+ (case Seq.list_of (Thm.bicompose false (false, rl, Thm.nprems_of rl) 1 revcut_rl') of
[th] => th
| _ => raise THM ("make_elim_preserve", 1, [rl]))
end;
--- a/src/Pure/Proof/proof_syntax.ML Sun Jul 05 21:10:23 2009 +0200
+++ b/src/Pure/Proof/proof_syntax.ML Thu Jul 09 08:55:42 2009 +0200
@@ -54,7 +54,7 @@
(Binding.name "AbsP", [propT, proofT --> proofT] ---> proofT, NoSyn),
(Binding.name "Hyp", propT --> proofT, NoSyn),
(Binding.name "Oracle", propT --> proofT, NoSyn),
- (Binding.name "Inclass", (Term.a_itselfT --> propT) --> proofT, NoSyn),
+ (Binding.name "OfClass", (Term.a_itselfT --> propT) --> proofT, NoSyn),
(Binding.name "MinProof", proofT, Delimfix "?")]
|> Sign.add_nonterminals [Binding.name "param", Binding.name "params"]
|> Sign.add_syntax_i
@@ -113,11 +113,11 @@
| NONE => error ("Unknown theorem " ^ quote name))
end
| _ => error ("Illegal proof constant name: " ^ quote s))
- | prf_of Ts (Const ("Inclass", _) $ Const (c_class, _)) =
+ | prf_of Ts (Const ("OfClass", _) $ Const (c_class, _)) =
(case try Logic.class_of_const c_class of
SOME c =>
change_type (if ty then SOME Ts else NONE)
- (Inclass (TVar ((Name.aT, 0), []), c))
+ (OfClass (TVar ((Name.aT, 0), []), c))
| NONE => error ("Bad class constant: " ^ quote c_class))
| prf_of Ts (Const ("Hyp", _) $ prop) = Hyp prop
| prf_of Ts (v as Var ((_, Type ("proof", _)))) = Hyp v
@@ -148,7 +148,7 @@
val AppPt = Const ("AppP", [proofT, proofT] ---> proofT);
val Hypt = Const ("Hyp", propT --> proofT);
val Oraclet = Const ("Oracle", propT --> proofT);
-val Inclasst = Const ("Inclass", (Term.itselfT dummyT --> propT) --> proofT);
+val OfClasst = Const ("OfClass", (Term.itselfT dummyT --> propT) --> proofT);
val MinProoft = Const ("MinProof", proofT);
val mk_tyapp = fold (fn T => fn prf => Const ("Appt",
@@ -161,8 +161,8 @@
| term_of _ (PAxm (name, _, NONE)) = Const (Long_Name.append "axm" name, proofT)
| term_of _ (PAxm (name, _, SOME Ts)) =
mk_tyapp Ts (Const (Long_Name.append "axm" name, proofT))
- | term_of _ (Inclass (T, c)) =
- mk_tyapp [T] (Inclasst $ Const (Logic.const_of_class c, Term.itselfT dummyT --> propT))
+ | term_of _ (OfClass (T, c)) =
+ mk_tyapp [T] (OfClasst $ Const (Logic.const_of_class c, Term.itselfT dummyT --> propT))
| term_of _ (PBound i) = Bound i
| term_of Ts (Abst (s, opT, prf)) =
let val T = the_default dummyT opT
--- a/src/Pure/Proof/reconstruct.ML Sun Jul 05 21:10:23 2009 +0200
+++ b/src/Pure/Proof/reconstruct.ML Thu Jul 09 08:55:42 2009 +0200
@@ -223,8 +223,8 @@
mk_cnstrts_atom env vTs prop opTs prf
| mk_cnstrts env _ _ vTs (prf as PAxm (_, prop, opTs)) =
mk_cnstrts_atom env vTs prop opTs prf
- | mk_cnstrts env _ _ vTs (prf as Inclass (T, c)) =
- mk_cnstrts_atom env vTs (Logic.mk_inclass (T, c)) NONE prf
+ | mk_cnstrts env _ _ vTs (prf as OfClass (T, c)) =
+ mk_cnstrts_atom env vTs (Logic.mk_of_class (T, c)) NONE prf
| mk_cnstrts env _ _ vTs (prf as Oracle (_, prop, opTs)) =
mk_cnstrts_atom env vTs prop opTs prf
| mk_cnstrts env _ _ vTs (Hyp t) = (t, Hyp t, [], env, vTs)
@@ -321,7 +321,7 @@
| prop_of0 Hs (Hyp t) = t
| prop_of0 Hs (PThm (_, ((_, prop, SOME Ts), _))) = prop_of_atom prop Ts
| prop_of0 Hs (PAxm (_, prop, SOME Ts)) = prop_of_atom prop Ts
- | prop_of0 Hs (Inclass (T, c)) = Logic.mk_inclass (T, c)
+ | prop_of0 Hs (OfClass (T, c)) = Logic.mk_of_class (T, c)
| prop_of0 Hs (Oracle (_, prop, SOME Ts)) = prop_of_atom prop Ts
| prop_of0 _ _ = error "prop_of: partial proof object";
--- a/src/Pure/axclass.ML Sun Jul 05 21:10:23 2009 +0200
+++ b/src/Pure/axclass.ML Thu Jul 09 08:55:42 2009 +0200
@@ -214,7 +214,7 @@
|> snd))
end;
-fun complete_arities thy =
+fun complete_arities thy =
let
val arities = snd (get_instances thy);
val (completions, arities') = arities
@@ -338,10 +338,11 @@
(* primitive rules *)
-fun add_classrel th thy =
+fun add_classrel raw_th thy =
let
+ val th = Thm.strip_shyps (Thm.transfer thy raw_th);
+ val prop = Thm.plain_prop_of th;
fun err () = raise THM ("add_classrel: malformed class relation", 0, [th]);
- val prop = Thm.plain_prop_of (Thm.transfer thy th);
val rel = Logic.dest_classrel prop handle TERM _ => err ();
val (c1, c2) = cert_classrel thy rel handle TYPE _ => err ();
in
@@ -351,10 +352,11 @@
|> perhaps complete_arities
end;
-fun add_arity th thy =
+fun add_arity raw_th thy =
let
+ val th = Thm.strip_shyps (Thm.transfer thy raw_th);
+ val prop = Thm.plain_prop_of th;
fun err () = raise THM ("add_arity: malformed type arity", 0, [th]);
- val prop = Thm.plain_prop_of (Thm.transfer thy th);
val (t, Ss, c) = Logic.dest_arity prop handle TERM _ => err ();
val T = Type (t, map TFree (Name.names Name.context Name.aT Ss));
val missing_params = Sign.complete_sort thy [c]
@@ -470,9 +472,9 @@
(* definition *)
- val conjs = map (curry Logic.mk_inclass (Term.aT [])) super @ flat axiomss;
+ val conjs = map (curry Logic.mk_of_class (Term.aT [])) super @ flat axiomss;
val class_eq =
- Logic.mk_equals (Logic.mk_inclass (Term.aT [], class), Logic.mk_conjunction_balanced conjs);
+ Logic.mk_equals (Logic.mk_of_class (Term.aT [], class), Logic.mk_conjunction_balanced conjs);
val ([def], def_thy) =
thy
@@ -575,12 +577,13 @@
fun derive_type _ (_, []) = []
| derive_type thy (typ, sort) =
let
+ val certT = Thm.ctyp_of thy;
val vars = Term.fold_atyps
(fn T as TFree (_, S) => insert (eq_fst op =) (T, S)
| T as TVar (_, S) => insert (eq_fst op =) (T, S)
| _ => I) typ [];
val hyps = vars
- |> map (fn (T, S) => (T, Thm.sort_triv thy (T, S) ~~ S));
+ |> map (fn (T, S) => (T, Thm.of_sort (certT T, S) ~~ S));
val ths = (typ, sort)
|> Sorts.of_sort_derivation (Syntax.pp_global thy) (Sign.classes_of thy)
{class_relation =
@@ -605,7 +608,7 @@
val ths =
sort |> map (fn c =>
Goal.finish (the (lookup_type cache' typ c) RS
- Goal.init (Thm.cterm_of thy (Logic.mk_inclass (typ, c))))
+ Goal.init (Thm.cterm_of thy (Logic.mk_of_class (typ, c))))
|> Thm.adjust_maxidx_thm ~1);
in (ths, cache') end;
--- a/src/Pure/drule.ML Sun Jul 05 21:10:23 2009 +0200
+++ b/src/Pure/drule.ML Thu Jul 09 08:55:42 2009 +0200
@@ -373,23 +373,25 @@
(*Rotates a rule's premises to the left by k*)
fun rotate_prems 0 = I
- | rotate_prems k = permute_prems 0 k;
+ | rotate_prems k = Thm.permute_prems 0 k;
fun with_subgoal i f = rotate_prems (i - 1) #> f #> rotate_prems (1 - i);
-(* permute prems, where the i-th position in the argument list (counting from 0)
- gives the position within the original thm to be transferred to position i.
- Any remaining trailing positions are left unchanged. *)
-val rearrange_prems = let
- fun rearr new [] thm = thm
- | rearr new (p::ps) thm = rearr (new+1)
- (map (fn q => if new<=q andalso q<p then q+1 else q) ps)
- (permute_prems (new+1) (new-p) (permute_prems new (p-new) thm))
+(*Permute prems, where the i-th position in the argument list (counting from 0)
+ gives the position within the original thm to be transferred to position i.
+ Any remaining trailing positions are left unchanged.*)
+val rearrange_prems =
+ let
+ fun rearr new [] thm = thm
+ | rearr new (p :: ps) thm =
+ rearr (new + 1)
+ (map (fn q => if new <= q andalso q < p then q + 1 else q) ps)
+ (Thm.permute_prems (new + 1) (new - p) (Thm.permute_prems new (p - new) thm))
in rearr 0 end;
(*Resolution: exactly one resolvent must be produced.*)
fun tha RSN (i,thb) =
- case Seq.chop 2 (biresolution false [(false,tha)] i thb) of
+ case Seq.chop 2 (Thm.biresolution false [(false,tha)] i thb) of
([th],_) => th
| ([],_) => raise THM("RSN: no unifiers", i, [tha,thb])
| _ => raise THM("RSN: multiple unifiers", i, [tha,thb]);
@@ -399,7 +401,7 @@
(*For joining lists of rules*)
fun thas RLN (i,thbs) =
- let val resolve = biresolution false (map (pair false) thas) i
+ let val resolve = Thm.biresolution false (map (pair false) thas) i
fun resb thb = Seq.list_of (resolve thb) handle THM _ => []
in maps resb thbs end;
@@ -425,7 +427,7 @@
with no lifting or renaming! Q may contain ==> or meta-quants
ALWAYS deletes premise i *)
fun compose(tha,i,thb) =
- distinct Thm.eq_thm (Seq.list_of (bicompose false (false,tha,0) i thb));
+ distinct Thm.eq_thm (Seq.list_of (Thm.bicompose false (false,tha,0) i thb));
fun compose_single (tha,i,thb) =
case compose (tha,i,thb) of
--- a/src/Pure/logic.ML Sun Jul 05 21:10:23 2009 +0200
+++ b/src/Pure/logic.ML Thu Jul 09 08:55:42 2009 +0200
@@ -36,8 +36,8 @@
val type_map: (term -> term) -> typ -> typ
val const_of_class: class -> string
val class_of_const: string -> class
- val mk_inclass: typ * class -> term
- val dest_inclass: term -> typ * class
+ val mk_of_class: typ * class -> term
+ val dest_of_class: term -> typ * class
val name_classrel: string * string -> string
val mk_classrel: class * class -> term
val dest_classrel: term -> class * class
@@ -219,13 +219,13 @@
handle Fail _ => raise TERM ("class_of_const: bad name " ^ quote c, []);
-(* class constraints *)
+(* class membership *)
-fun mk_inclass (ty, c) =
+fun mk_of_class (ty, c) =
Const (const_of_class c, Term.itselfT ty --> propT) $ mk_type ty;
-fun dest_inclass (t as Const (c_class, _) $ ty) = (dest_type ty, class_of_const c_class)
- | dest_inclass t = raise TERM ("dest_inclass", [t]);
+fun dest_of_class (t as Const (c_class, _) $ ty) = (dest_type ty, class_of_const c_class)
+ | dest_of_class t = raise TERM ("dest_of_class", [t]);
(* class relations *)
@@ -233,10 +233,10 @@
fun name_classrel (c1, c2) =
Long_Name.base_name c1 ^ "_" ^ Long_Name.base_name c2;
-fun mk_classrel (c1, c2) = mk_inclass (Term.aT [c1], c2);
+fun mk_classrel (c1, c2) = mk_of_class (Term.aT [c1], c2);
fun dest_classrel tm =
- (case dest_inclass tm of
+ (case dest_of_class tm of
(TVar (_, [c1]), c2) => (c1, c2)
| _ => raise TERM ("dest_classrel", [tm]));
@@ -251,13 +251,13 @@
fun mk_arities (t, Ss, S) =
let val T = Type (t, ListPair.map TFree (Name.invents Name.context Name.aT (length Ss), Ss))
- in map (fn c => mk_inclass (T, c)) S end;
+ in map (fn c => mk_of_class (T, c)) S end;
fun dest_arity tm =
let
fun err () = raise TERM ("dest_arity", [tm]);
- val (ty, c) = dest_inclass tm;
+ val (ty, c) = dest_of_class tm;
val (t, tvars) =
(case ty of
Type (t, tys) => (t, map dest_TVar tys handle TYPE _ => err ())
--- a/src/Pure/more_thm.ML Sun Jul 05 21:10:23 2009 +0200
+++ b/src/Pure/more_thm.ML Thu Jul 09 08:55:42 2009 +0200
@@ -26,11 +26,11 @@
val eq_thm_thy: thm * thm -> bool
val eq_thm_prop: thm * thm -> bool
val equiv_thm: thm * thm -> bool
- val sort_triv: theory -> typ * sort -> thm list
+ val class_triv: theory -> class -> thm
+ val of_sort: ctyp * sort -> thm list
val check_shyps: sort list -> thm -> thm
val is_dummy: thm -> bool
val plain_prop_of: thm -> term
- val fold_terms: (term -> 'a -> 'a) -> thm -> 'a -> 'a
val add_thm: thm -> thm list -> thm list
val del_thm: thm -> thm list -> thm list
val merge_thms: thm list * thm list -> thm list
@@ -171,14 +171,10 @@
(* type classes and sorts *)
-fun sort_triv thy (T, S) =
- let
- val certT = Thm.ctyp_of thy;
- val cT = certT T;
- fun class_triv c =
- Thm.class_triv thy c
- |> Thm.instantiate ([(certT (TVar ((Name.aT, 0), [c])), cT)], []);
- in map class_triv S end;
+fun class_triv thy c =
+ Thm.of_class (Thm.ctyp_of thy (TVar ((Name.aT, 0), [c])), c);
+
+fun of_sort (T, S) = map (fn c => Thm.of_class (T, c)) S;
fun check_shyps sorts raw_th =
let
@@ -214,10 +210,6 @@
else prop
end;
-fun fold_terms f th =
- let val {tpairs, prop, hyps, ...} = Thm.rep_thm th
- in fold (fn (t, u) => f t #> f u) tpairs #> f prop #> fold f hyps end;
-
(* collections of theorems in canonical order *)
--- a/src/Pure/proofterm.ML Sun Jul 05 21:10:23 2009 +0200
+++ b/src/Pure/proofterm.ML Thu Jul 09 08:55:42 2009 +0200
@@ -19,7 +19,7 @@
| op %% of proof * proof
| Hyp of term
| PAxm of string * term * typ list option
- | Inclass of typ * class
+ | OfClass of typ * class
| Oracle of string * term * typ list option
| Promise of serial * term * typ list
| PThm of serial * ((string * term * typ list option) * proof_body future)
@@ -141,7 +141,7 @@
| op %% of proof * proof
| Hyp of term
| PAxm of string * term * typ list option
- | Inclass of typ * class
+ | OfClass of typ * class
| Oracle of string * term * typ list option
| Promise of serial * term * typ list
| PThm of serial * ((string * term * typ list option) * proof_body future)
@@ -292,7 +292,7 @@
| mapp (prf1 %% prf2) = (mapp prf1 %% mapph prf2
handle SAME => prf1 %% mapp prf2)
| mapp (PAxm (a, prop, SOME Ts)) = PAxm (a, prop, SOME (map_typs Ts))
- | mapp (Inclass (T, c)) = Inclass (apsome g (SOME T) |> the, c)
+ | mapp (OfClass (T, c)) = OfClass (apsome g (SOME T) |> the, c)
| mapp (Oracle (a, prop, SOME Ts)) = Oracle (a, prop, SOME (map_typs Ts))
| mapp (Promise (i, prop, Ts)) = Promise (i, prop, map_typs Ts)
| mapp (PThm (i, ((a, prop, SOME Ts), body))) =
@@ -320,7 +320,7 @@
| fold_proof_terms f g (prf1 %% prf2) =
fold_proof_terms f g prf1 #> fold_proof_terms f g prf2
| fold_proof_terms _ g (PAxm (_, _, SOME Ts)) = fold g Ts
- | fold_proof_terms _ g (Inclass (T, _)) = g T
+ | fold_proof_terms _ g (OfClass (T, _)) = g T
| fold_proof_terms _ g (Oracle (_, _, SOME Ts)) = fold g Ts
| fold_proof_terms _ g (Promise (_, _, Ts)) = fold g Ts
| fold_proof_terms _ g (PThm (_, ((_, _, SOME Ts), _))) = fold g Ts
@@ -335,7 +335,7 @@
| size_of_proof _ = 1;
fun change_type opTs (PAxm (name, prop, _)) = PAxm (name, prop, opTs)
- | change_type (SOME [T]) (Inclass (_, c)) = Inclass (T, c)
+ | change_type (SOME [T]) (OfClass (_, c)) = OfClass (T, c)
| change_type opTs (Oracle (name, prop, _)) = Oracle (name, prop, opTs)
| change_type opTs (Promise _) = error "change_type: unexpected promise"
| change_type opTs (PThm (i, ((name, prop, _), body))) = PThm (i, ((name, prop, opTs), body))
@@ -473,7 +473,7 @@
| norm (prf1 %% prf2) = (norm prf1 %% normh prf2
handle SAME => prf1 %% norm prf2)
| norm (PAxm (s, prop, Ts)) = PAxm (s, prop, apsome' (htypeTs norm_types_same) Ts)
- | norm (Inclass (T, c)) = Inclass (htypeT norm_type_same T, c)
+ | norm (OfClass (T, c)) = OfClass (htypeT norm_type_same T, c)
| norm (Oracle (s, prop, Ts)) = Oracle (s, prop, apsome' (htypeTs norm_types_same) Ts)
| norm (Promise (i, prop, Ts)) = Promise (i, prop, htypeTs norm_types_same Ts)
| norm (PThm (i, ((s, t, Ts), body))) =
@@ -719,8 +719,8 @@
handle SAME => prf1 %% lift' Us Ts prf2)
| lift' _ _ (PAxm (s, prop, Ts)) =
PAxm (s, prop, apsome' (same (op =) (map (Logic.incr_tvar inc))) Ts)
- | lift' _ _ (Inclass (T, c)) =
- Inclass (same (op =) (Logic.incr_tvar inc) T, c)
+ | lift' _ _ (OfClass (T, c)) =
+ OfClass (same (op =) (Logic.incr_tvar inc) T, c)
| lift' _ _ (Oracle (s, prop, Ts)) =
Oracle (s, prop, apsome' (same (op =) (map (Logic.incr_tvar inc))) Ts)
| lift' _ _ (Promise (i, prop, Ts)) =
@@ -977,7 +977,7 @@
orelse exists #1 prfs then [i] else [], false, map (pair false) ts, prf)
| shrink' ls lev ts prfs (prf as Hyp _) = ([], false, map (pair false) ts, prf)
| shrink' ls lev ts prfs (prf as MinProof) = ([], false, map (pair false) ts, prf)
- | shrink' ls lev ts prfs (prf as Inclass _) = ([], false, map (pair false) ts, prf)
+ | shrink' ls lev ts prfs (prf as OfClass _) = ([], false, map (pair false) ts, prf)
| shrink' ls lev ts prfs prf =
let
val prop =
@@ -1069,7 +1069,7 @@
| mtch Ts i j inst (PAxm (s1, _, opTs), PAxm (s2, _, opUs)) =
if s1 = s2 then optmatch matchTs inst (opTs, opUs)
else raise PMatch
- | mtch Ts i j inst (Inclass (T1, c1), Inclass (T2, c2)) =
+ | mtch Ts i j inst (OfClass (T1, c1), OfClass (T2, c2)) =
if c1 = c2 then matchT inst (T1, T2)
else raise PMatch
| mtch Ts i j inst (PThm (_, ((name1, prop1, opTs), _)), PThm (_, ((name2, prop2, opUs), _))) =
@@ -1103,7 +1103,7 @@
NONE => prf
| SOME prf' => incr_pboundvars plev tlev prf')
| subst _ _ (PAxm (id, prop, Ts)) = PAxm (id, prop, Option.map (map substT) Ts)
- | subst _ _ (Inclass (T, c)) = Inclass (substT T, c)
+ | subst _ _ (OfClass (T, c)) = OfClass (substT T, c)
| subst _ _ (Oracle (id, prop, Ts)) = Oracle (id, prop, Option.map (map substT) Ts)
| subst _ _ (Promise (i, prop, Ts)) = Promise (i, prop, (map substT) Ts)
| subst _ _ (PThm (i, ((id, prop, Ts), body))) =
@@ -1130,7 +1130,7 @@
(_, Hyp (Var _)) => true
| (Hyp (Var _), _) => true
| (PAxm (a, _, _), PAxm (b, _, _)) => a = b andalso matchrands prf1 prf2
- | (Inclass (_, c), Inclass (_, d)) => c = d andalso matchrands prf1 prf2
+ | (OfClass (_, c), OfClass (_, d)) => c = d andalso matchrands prf1 prf2
| (PThm (_, ((a, propa, _), _)), PThm (_, ((b, propb, _), _))) =>
a = b andalso propa = propb andalso matchrands prf1 prf2
| (PBound i, PBound j) => i = j andalso matchrands prf1 prf2
--- a/src/Pure/sign.ML Sun Jul 05 21:10:23 2009 +0200
+++ b/src/Pure/sign.ML Thu Jul 09 08:55:42 2009 +0200
@@ -27,7 +27,7 @@
val defaultS: theory -> sort
val subsort: theory -> sort * sort -> bool
val of_sort: theory -> typ * sort -> bool
- val witness_sorts: theory -> sort list -> sort list -> (typ * sort) list
+ val witness_sorts: theory -> (typ * sort) list -> sort list -> (typ * sort) list
val is_logtype: theory -> string -> bool
val typ_instance: theory -> typ * typ -> bool
val typ_equiv: theory -> typ * typ -> bool
--- a/src/Pure/sorts.ML Sun Jul 05 21:10:23 2009 +0200
+++ b/src/Pure/sorts.ML Thu Jul 09 08:55:42 2009 +0200
@@ -62,7 +62,7 @@
type_constructor: string -> ('a * class) list list -> class -> 'a,
type_variable: typ -> ('a * class) list} ->
typ * sort -> 'a list (*exception CLASS_ERROR*)
- val witness_sorts: algebra -> string list -> sort list -> sort list -> (typ * sort) list
+ val witness_sorts: algebra -> string list -> (typ * sort) list -> sort list -> (typ * sort) list
end;
structure Sorts: SORTS =
@@ -432,18 +432,17 @@
fun witness_sorts algebra types hyps sorts =
let
fun le S1 S2 = sort_le algebra (S1, S2);
- fun get_solved S2 (T, S1) = if le S1 S2 then SOME (T, S2) else NONE;
- fun get_hyp S2 S1 = if le S1 S2 then SOME (TFree ("'hyp", S1), S2) else NONE;
+ fun get S2 (T, S1) = if le S1 S2 then SOME (T, S2) else NONE;
fun mg_dom t S = SOME (mg_domain algebra t S) handle CLASS_ERROR _ => NONE;
fun witn_sort _ [] solved_failed = (SOME (propT, []), solved_failed)
| witn_sort path S (solved, failed) =
if exists (le S) failed then (NONE, (solved, failed))
else
- (case get_first (get_solved S) solved of
+ (case get_first (get S) solved of
SOME w => (SOME w, (solved, failed))
| NONE =>
- (case get_first (get_hyp S) hyps of
+ (case get_first (get S) hyps of
SOME w => (SOME w, (w :: solved, failed))
| NONE => witn_types path types S (solved, failed)))
--- a/src/Pure/tactic.ML Sun Jul 05 21:10:23 2009 +0200
+++ b/src/Pure/tactic.ML Thu Jul 09 08:55:42 2009 +0200
@@ -105,24 +105,24 @@
thus (REPEAT (resolve_tac rules i)) stops once subgoal i disappears. *)
(*Solve subgoal i by assumption*)
-fun assume_tac i = PRIMSEQ (assumption i);
+fun assume_tac i = PRIMSEQ (Thm.assumption i);
(*Solve subgoal i by assumption, using no unification*)
-fun eq_assume_tac i = PRIMITIVE (eq_assumption i);
+fun eq_assume_tac i = PRIMITIVE (Thm.eq_assumption i);
(** Resolution/matching tactics **)
(*The composition rule/state: no lifting or var renaming.
- The arg = (bires_flg, orule, m) ; see bicompose for explanation.*)
-fun compose_tac arg i = PRIMSEQ (bicompose false arg i);
+ The arg = (bires_flg, orule, m); see Thm.bicompose for explanation.*)
+fun compose_tac arg i = PRIMSEQ (Thm.bicompose false arg i);
(*Converts a "destruct" rule like P&Q==>P to an "elimination" rule
like [| P&Q; P==>R |] ==> R *)
fun make_elim rl = zero_var_indexes (rl RS revcut_rl);
(*Attack subgoal i by resolution, using flags to indicate elimination rules*)
-fun biresolve_tac brules i = PRIMSEQ (biresolution false brules i);
+fun biresolve_tac brules i = PRIMSEQ (Thm.biresolution false brules i);
(*Resolution: the simple case, works for introduction rules*)
fun resolve_tac rules = biresolve_tac (map (pair false) rules);
@@ -152,7 +152,7 @@
fun solve_tac rules = resolve_tac rules THEN_ALL_NEW assume_tac;
(*Matching tactics -- as above, but forbid updating of state*)
-fun bimatch_tac brules i = PRIMSEQ (biresolution true brules i);
+fun bimatch_tac brules i = PRIMSEQ (Thm.biresolution true brules i);
fun match_tac rules = bimatch_tac (map (pair false) rules);
fun ematch_tac rules = bimatch_tac (map (pair true) rules);
fun dmatch_tac rls = ematch_tac (map make_elim rls);
@@ -295,7 +295,7 @@
let val hyps = Logic.strip_assums_hyp prem
and concl = Logic.strip_assums_concl prem
val kbrls = Net.unify_term inet concl @ maps (Net.unify_term enet) hyps
- in PRIMSEQ (biresolution match (order kbrls) i) end);
+ in PRIMSEQ (Thm.biresolution match (order kbrls) i) end);
(*versions taking pre-built nets. No filtering of brls*)
val biresolve_from_nets_tac = biresolution_from_nets_tac order_list false;
@@ -326,7 +326,7 @@
in
if pred krls
then PRIMSEQ
- (biresolution match (map (pair false) (order_list krls)) i)
+ (Thm.biresolution match (map (pair false) (order_list krls)) i)
else no_tac
end);
@@ -359,15 +359,15 @@
fun rename_tac xs i =
case Library.find_first (not o Syntax.is_identifier) xs of
SOME x => error ("Not an identifier: " ^ x)
- | NONE => PRIMITIVE (rename_params_rule (xs, i));
+ | NONE => PRIMITIVE (Thm.rename_params_rule (xs, i));
(*rotate_tac n i: rotate the assumptions of subgoal i by n positions, from
right to left if n is positive, and from left to right if n is negative.*)
fun rotate_tac 0 i = all_tac
- | rotate_tac k i = PRIMITIVE (rotate_rule k i);
+ | rotate_tac k i = PRIMITIVE (Thm.rotate_rule k i);
(*Rotates the given subgoal to be the last.*)
-fun defer_tac i = PRIMITIVE (permute_prems (i-1) 1);
+fun defer_tac i = PRIMITIVE (Thm.permute_prems (i - 1) 1);
(* remove premises that do not satisfy p; fails if all prems satisfy p *)
fun filter_prems_tac p =
--- a/src/Pure/tctical.ML Sun Jul 05 21:10:23 2009 +0200
+++ b/src/Pure/tctical.ML Thu Jul 09 08:55:42 2009 +0200
@@ -463,8 +463,7 @@
(forall_intr_list cparams (implies_intr_list chyps Cth)))
end
(*function to replace the current subgoal*)
- fun next st = bicompose false (false, relift st, nprems_of st)
- gno state
+ fun next st = Thm.bicompose false (false, relift st, nprems_of st) gno state
in Seq.maps next (tacf subprems (trivial (cterm concl))) end;
end;
--- a/src/Pure/thm.ML Sun Jul 05 21:10:23 2009 +0200
+++ b/src/Pure/thm.ML Thu Jul 09 08:55:42 2009 +0200
@@ -92,19 +92,11 @@
val instantiate: (ctyp * ctyp) list * (cterm * cterm) list -> thm -> thm
val instantiate_cterm: (ctyp * ctyp) list * (cterm * cterm) list -> cterm -> cterm
val trivial: cterm -> thm
- val class_triv: theory -> class -> thm
+ val of_class: ctyp * class -> thm
val unconstrainT: ctyp -> thm -> thm
val dest_state: thm * int -> (term * term) list * term list * term * term
val lift_rule: cterm -> thm -> thm
val incr_indexes: int -> thm -> thm
- val assumption: int -> thm -> thm Seq.seq
- val eq_assumption: int -> thm -> thm
- val rotate_rule: int -> int -> thm -> thm
- val permute_prems: int -> int -> thm -> thm
- val rename_params_rule: string list * int -> thm -> thm
- val compose_no_flatten: bool -> thm * int -> int -> thm -> thm Seq.seq
- val bicompose: bool -> bool * thm * int -> int -> thm -> thm Seq.seq
- val biresolution: bool -> (bool * thm) list -> int -> thm -> thm Seq.seq
end;
signature THM =
@@ -117,37 +109,46 @@
val dest_fun2: cterm -> cterm
val dest_arg1: cterm -> cterm
val dest_abs: string option -> cterm -> cterm * cterm
- val adjust_maxidx_cterm: int -> cterm -> cterm
val capply: cterm -> cterm -> cterm
val cabs: cterm -> cterm -> cterm
- val major_prem_of: thm -> term
- val no_prems: thm -> bool
+ val adjust_maxidx_cterm: int -> cterm -> cterm
+ val incr_indexes_cterm: int -> cterm -> cterm
+ val match: cterm * cterm -> (ctyp * ctyp) list * (cterm * cterm) list
+ val first_order_match: cterm * cterm -> (ctyp * ctyp) list * (cterm * cterm) list
+ val fold_terms: (term -> 'a -> 'a) -> thm -> 'a -> 'a
val terms_of_tpairs: (term * term) list -> term list
+ val full_prop_of: thm -> term
val maxidx_of: thm -> int
val maxidx_thm: thm -> int -> int
val hyps_of: thm -> term list
- val full_prop_of: thm -> term
+ val no_prems: thm -> bool
+ val major_prem_of: thm -> term
val axiom: theory -> string -> thm
val axioms_of: theory -> (string * thm) list
- val get_name: thm -> string
- val put_name: string -> thm -> thm
val get_tags: thm -> Properties.T
val map_tags: (Properties.T -> Properties.T) -> thm -> thm
val norm_proof: thm -> thm
val adjust_maxidx_thm: int -> thm -> thm
- val rename_boundvars: term -> term -> thm -> thm
- val match: cterm * cterm -> (ctyp * ctyp) list * (cterm * cterm) list
- val first_order_match: cterm * cterm -> (ctyp * ctyp) list * (cterm * cterm) list
- val incr_indexes_cterm: int -> cterm -> cterm
val varifyT: thm -> thm
val varifyT': (string * sort) list -> thm -> ((string * sort) * indexname) list * thm
val freezeT: thm -> thm
+ val assumption: int -> thm -> thm Seq.seq
+ val eq_assumption: int -> thm -> thm
+ val rotate_rule: int -> int -> thm -> thm
+ val permute_prems: int -> int -> thm -> thm
+ val rename_params_rule: string list * int -> thm -> thm
+ val compose_no_flatten: bool -> thm * int -> int -> thm -> thm Seq.seq
+ val bicompose: bool -> bool * thm * int -> int -> thm -> thm Seq.seq
+ val biresolution: bool -> (bool * thm) list -> int -> thm -> thm Seq.seq
+ val rename_boundvars: term -> term -> thm -> thm
val future: thm future -> cterm -> thm
val pending_groups: thm -> Task_Queue.group list -> Task_Queue.group list
val status_of: thm -> {oracle: bool, unfinished: bool, failed: bool}
val proof_body_of: thm -> proof_body
val proof_of: thm -> proof
val join_proof: thm -> unit
+ val get_name: thm -> string
+ val put_name: string -> thm -> thm
val extern_oracles: theory -> xstring list
val add_oracle: binding * ('a -> cterm) -> theory -> (string * ('a -> thm)) * theory
end;
@@ -367,6 +368,9 @@
prop = cterm maxidx prop}
end;
+fun fold_terms f (Thm (_, {tpairs, prop, hyps, ...})) =
+ fold (fn (t, u) => f t #> f u) tpairs #> f prop #> fold f hyps;
+
fun terms_of_tpairs tpairs = fold_rev (fn (t, u) => cons t o cons u) tpairs [];
fun eq_tpairs ((t, u), (t', u')) = t aconv t' andalso u aconv u';
@@ -473,29 +477,28 @@
(** sort contexts of theorems **)
-fun present_sorts (Thm (_, {hyps, tpairs, prop, ...})) =
- fold (fn (t, u) => Sorts.insert_term t o Sorts.insert_term u) tpairs
- (Sorts.insert_terms hyps (Sorts.insert_term prop []));
-
-(*remove extra sorts that are non-empty by virtue of type signature information*)
+(*remove extra sorts that are witnessed by type signature information*)
fun strip_shyps (thm as Thm (_, {shyps = [], ...})) = thm
| strip_shyps (thm as Thm (der, {thy_ref, tags, maxidx, shyps, hyps, tpairs, prop})) =
let
val thy = Theory.deref thy_ref;
- val present = present_sorts thm;
- val extra = Sorts.subtract present shyps;
- val extra' =
- Sorts.subtract (map #2 (Sign.witness_sorts thy present extra)) extra
+ val present =
+ (fold_terms o fold_types o fold_atyps)
+ (fn T as TFree (_, S) => insert (eq_snd op =) (T, S)
+ | T as TVar (_, S) => insert (eq_snd op =) (T, S)) thm [];
+ val extra = fold (Sorts.remove_sort o #2) present shyps;
+ val witnessed = Sign.witness_sorts thy present extra;
+ val extra' = fold (Sorts.remove_sort o #2) witnessed extra
|> Sorts.minimal_sorts (Sign.classes_of thy);
- val shyps' = Sorts.union present extra'
- |> Sorts.remove_sort [];
+ val shyps' = fold (Sorts.insert_sort o #2) present extra';
in
Thm (der, {thy_ref = Theory.check_thy thy, tags = tags, maxidx = maxidx,
shyps = shyps', hyps = hyps, tpairs = tpairs, prop = prop})
end;
(*dangling sort constraints of a thm*)
-fun extra_shyps (th as Thm (_, {shyps, ...})) = Sorts.subtract (present_sorts th) shyps;
+fun extra_shyps (th as Thm (_, {shyps, ...})) =
+ Sorts.subtract (fold_terms Sorts.insert_term th []) shyps;
@@ -1110,22 +1113,28 @@
tpairs = [],
prop = Logic.mk_implies (A, A)});
-(*Axiom-scheme reflecting signature contents: "OFCLASS(?'a::c, c_class)" *)
-fun class_triv thy raw_c =
+(*Axiom-scheme reflecting signature contents
+ T :: c
+ -------------------
+ OFCLASS(T, c_class)
+*)
+fun of_class (cT, raw_c) =
let
+ val Ctyp {thy_ref, T, ...} = cT;
+ val thy = Theory.deref thy_ref;
val c = Sign.certify_class thy raw_c;
- val T = TVar ((Name.aT, 0), [c]);
- val Cterm {t = prop, maxidx, sorts, ...} = cterm_of thy (Logic.mk_inclass (T, c))
- handle TERM (msg, _) => raise THM ("class_triv: " ^ msg, 0, []);
+ val Cterm {t = prop, maxidx, sorts, ...} = cterm_of thy (Logic.mk_of_class (T, c));
in
- Thm (deriv_rule0 (Pt.Inclass (T, c)),
- {thy_ref = Theory.check_thy thy,
- tags = [],
- maxidx = maxidx,
- shyps = sorts,
- hyps = [],
- tpairs = [],
- prop = prop})
+ if Sign.of_sort thy (T, [c]) then
+ Thm (deriv_rule0 (Pt.OfClass (T, c)),
+ {thy_ref = Theory.check_thy thy,
+ tags = [],
+ maxidx = maxidx,
+ shyps = sorts,
+ hyps = [],
+ tpairs = [],
+ prop = prop})
+ else raise THM ("of_class: type not of class " ^ Syntax.string_of_sort_global thy [c], 0, [])
end;
(*Internalize sort constraints of type variable*)
@@ -1137,7 +1146,7 @@
raise THM ("unconstrainT: not a type variable", 0, [th]);
val T' = TVar ((x, i), []);
val unconstrain = Term.map_types (Term.map_atyps (fn U => if U = T then T' else U));
- val constraints = map (curry Logic.mk_inclass T') S;
+ val constraints = map (curry Logic.mk_of_class T') S;
in
Thm (deriv_rule0 (Pt.PAxm ("Pure.unconstrainT", prop, SOME [])),
{thy_ref = Theory.merge_refs (thy_ref1, thy_ref2),
--- a/src/Pure/type.ML Sun Jul 05 21:10:23 2009 +0200
+++ b/src/Pure/type.ML Thu Jul 09 08:55:42 2009 +0200
@@ -27,7 +27,7 @@
val inter_sort: tsig -> sort * sort -> sort
val cert_class: tsig -> class -> class
val cert_sort: tsig -> sort -> sort
- val witness_sorts: tsig -> sort list -> sort list -> (typ * sort) list
+ val witness_sorts: tsig -> (typ * sort) list -> sort list -> (typ * sort) list
type mode
val mode_default: mode
val mode_syntax: mode
--- a/src/Tools/Code/code_preproc.ML Sun Jul 05 21:10:23 2009 +0200
+++ b/src/Tools/Code/code_preproc.ML Thu Jul 09 08:55:42 2009 +0200
@@ -102,6 +102,15 @@
fun rhs_conv conv thm = Thm.transitive thm ((conv o Thm.rhs_of) thm);
+fun eqn_conv conv =
+ let
+ fun lhs_conv ct = if can Thm.dest_comb ct
+ then Conv.combination_conv lhs_conv conv ct
+ else Conv.all_conv ct;
+ in Conv.combination_conv (Conv.arg_conv lhs_conv) conv end;
+
+val rewrite_eqn = Conv.fconv_rule o eqn_conv o Simplifier.rewrite;
+
fun term_of_conv thy f =
Thm.cterm_of thy
#> f
@@ -117,7 +126,7 @@
in
eqns
|> apply_functrans thy c functrans
- |> (map o apfst) (Code.rewrite_eqn pre)
+ |> (map o apfst) (rewrite_eqn pre)
|> (map o apfst) (AxClass.unoverload thy)
|> map (Code.assert_eqn thy)
|> burrow_fst (Code.norm_args thy)
@@ -213,9 +222,19 @@
(fn Const (c, ty) => insert (op =) (c, Sign.const_typargs thy (c, Logic.unvarifyT ty)) | _ => I)
(map (op :: o swap o apfst (snd o strip_comb) o Logic.dest_equals o Thm.plain_prop_of o fst) eqns);
+fun default_typscheme_of thy c =
+ let
+ val ty = (snd o dest_Const o TermSubst.zero_var_indexes o curry Const c
+ o Type.strip_sorts o Sign.the_const_type thy) c;
+ in case AxClass.class_of_param thy c
+ of SOME class => ([(Name.aT, [class])], ty)
+ | NONE => Code.typscheme thy (c, ty)
+ end;
+
fun tyscm_rhss_of thy c eqns =
let
- val tyscm = case eqns of [] => Code.default_typscheme thy c
+ val tyscm = case eqns
+ of [] => default_typscheme_of thy c
| ((thm, _) :: _) => Code.typscheme_eqn thy thm;
val rhss = consts_of thy eqns;
in (tyscm, rhss) end;
@@ -381,6 +400,17 @@
handle Sorts.CLASS_ERROR _ => [] (*permissive!*))
end;
+fun inst_thm thy tvars' thm =
+ let
+ val tvars = (Term.add_tvars o Thm.prop_of) thm [];
+ val inter_sort = Sorts.inter_sort (Sign.classes_of thy);
+ fun mk_inst (tvar as (v, sort)) = case Vartab.lookup tvars' v
+ of SOME sort' => SOME (pairself (Thm.ctyp_of thy o TVar)
+ (tvar, (v, inter_sort (sort, sort'))))
+ | NONE => NONE;
+ val insts = map_filter mk_inst tvars;
+ in Thm.instantiate (insts, []) thm end;
+
fun add_arity thy vardeps (class, tyco) =
AList.default (op =)
((class, tyco), map (fn k => (snd o Vargraph.get_node vardeps) (Inst (class, tyco), k))
@@ -394,7 +424,7 @@
val inst_tab = Vartab.empty |> fold (fn (v, sort) =>
Vartab.update ((v, 0), sort)) lhs;
val eqns = proto_eqns
- |> (map o apfst) (Code.inst_thm thy inst_tab);
+ |> (map o apfst) (inst_thm thy inst_tab);
val (tyscm, rhss') = tyscm_rhss_of thy c eqns;
val eqngr' = Graph.new_node (c, (tyscm, eqns)) eqngr;
in (map (pair c) rhss' @ rhss, eqngr') end;
@@ -420,7 +450,7 @@
(** store for preprocessed arities and code equations **)
-structure Wellsorted = CodeDataFun
+structure Wellsorted = Code_Data_Fun
(
type T = ((string * class) * sort list) list * code_graph;
val empty = ([], Graph.empty);
--- a/src/Tools/Code/code_target.ML Sun Jul 05 21:10:23 2009 +0200
+++ b/src/Tools/Code/code_target.ML Thu Jul 09 08:55:42 2009 +0200
@@ -286,7 +286,7 @@
fun gen_add_syntax_const prep_const target raw_c raw_syn thy =
let
val c = prep_const thy raw_c;
- fun check_args (syntax as (n, _)) = if n > Code.no_args thy c
+ fun check_args (syntax as (n, _)) = if n > Code.args_number thy c
then error ("Too many arguments in syntax for constant " ^ quote c)
else syntax;
in case raw_syn
--- a/src/Tools/Code/code_thingol.ML Sun Jul 05 21:10:23 2009 +0200
+++ b/src/Tools/Code/code_thingol.ML Thu Jul 09 08:55:42 2009 +0200
@@ -627,8 +627,8 @@
fun arg_types num_args ty = (fst o chop num_args o fst o strip_type) ty;
val tys = arg_types num_args (snd c_ty);
val ty = nth tys t_pos;
- fun mk_constr c t = let val n = Code.no_args thy c
- in ((c, arg_types (Code.no_args thy c) (fastype_of t) ---> ty), n) end;
+ fun mk_constr c t = let val n = Code.args_number thy c
+ in ((c, arg_types n (fastype_of t) ---> ty), n) end;
val constrs = if null case_pats then []
else map2 mk_constr case_pats (nth_drop t_pos ts);
fun casify naming constrs ty ts =
@@ -722,15 +722,15 @@
ensure_inst thy algbr funcgr inst
##>> (fold_map o fold_map) mk_dict yss
#>> (fn (inst, dss) => DictConst (inst, dss))
- | mk_dict (Local (classrels, (v, (k, sort)))) =
+ | mk_dict (Local (classrels, (v, (n, sort)))) =
fold_map (ensure_classrel thy algbr funcgr) classrels
- #>> (fn classrels => DictVar (classrels, (unprefix "'" v, (k, length sort))))
+ #>> (fn classrels => DictVar (classrels, (unprefix "'" v, (n, length sort))))
in fold_map mk_dict typargs end;
(* store *)
-structure Program = CodeDataFun
+structure Program = Code_Data_Fun
(
type T = naming * program;
val empty = (empty_naming, Graph.empty);
--- a/src/Tools/IsaPlanner/isand.ML Sun Jul 05 21:10:23 2009 +0200
+++ b/src/Tools/IsaPlanner/isand.ML Thu Jul 09 08:55:42 2009 +0200
@@ -110,7 +110,7 @@
fun solve_with sol th =
let fun solvei 0 = Seq.empty
| solvei i =
- Seq.append (bicompose false (false,sol,0) i th) (solvei (i - 1))
+ Seq.append (Thm.bicompose false (false,sol,0) i th) (solvei (i - 1))
in
solvei (Thm.nprems_of th)
end;
@@ -337,7 +337,7 @@
val newth' = Drule.implies_intr_list sgallcts allified_newth
in
- bicompose false (false, newth', (length sgallcts)) i gth
+ Thm.bicompose false (false, newth', (length sgallcts)) i gth
end;
(*
--- a/src/Tools/nbe.ML Sun Jul 05 21:10:23 2009 +0200
+++ b/src/Tools/nbe.ML Thu Jul 09 08:55:42 2009 +0200
@@ -393,10 +393,11 @@
let
val ts' = take_until is_dict ts;
val c = const_of_idx idx;
- val (_, T) = Code.default_typscheme thy c;
- val T' = map_type_tfree (fn (v, _) => TypeInfer.param typidx (v, [])) T;
+ val T = map_type_tvar (fn ((v, i), _) =>
+ TypeInfer.param typidx (v ^ string_of_int i, []))
+ (Sign.the_const_type thy c);
val typidx' = typidx + 1;
- in of_apps bounds (Term.Const (c, T'), ts') typidx' end
+ in of_apps bounds (Term.Const (c, T), ts') typidx' end
| of_univ bounds (BVar (n, ts)) typidx =
of_apps bounds (Bound (bounds - n - 1), ts) typidx
| of_univ bounds (t as Abs _) typidx =
@@ -407,7 +408,7 @@
(* function store *)
-structure Nbe_Functions = CodeDataFun
+structure Nbe_Functions = Code_Data_Fun
(
type T = Code_Thingol.naming * ((Univ option * int) Graph.T * (int * string Inttab.table));
val empty = (Code_Thingol.empty_naming, (Graph.empty, (0, Inttab.empty)));