sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
"num" syntax (still with "#"), Numeral0, Numeral1;
--- a/src/HOL/Algebra/poly/PolyHomo.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Algebra/poly/PolyHomo.ML Fri Oct 05 21:52:39 2001 +0200
@@ -112,15 +112,15 @@
(* Examples *)
Goal
- "EVAL (x::'a::domain) (a*X^2 + b*X^1 + c*X^0) = a * x ^ 2 + b * x ^ 1 + c";
+ "EVAL (x::'a::domain) (a*X^# 2 + b*X^1 + c*X^0) = a * x ^ # 2 + b * x ^ 1 + c";
by (asm_simp_tac (simpset() delsimps [power_Suc]
addsimps [EVAL_homo, EVAL_monom, EVAL_smult]) 1);
result();
Goal
"EVAL (y::'a::domain) \
-\ (EVAL (const x) (monom 1 + const (a*X^2 + b*X^1 + c*X^0))) = \
-\ x ^ 1 + (a * y ^ 2 + b * y ^ 1 + c)";
+\ (EVAL (const x) (monom 1 + const (a*X^# 2 + b*X^1 + c*X^0))) = \
+\ x ^ 1 + (a * y ^ # 2 + b * y ^ 1 + c)";
by (asm_simp_tac (simpset() delsimps [power_Suc]
addsimps [EVAL_homo, EVAL_monom, EVAL_smult, EVAL_const]) 1);
result();
--- a/src/HOL/Auth/KerberosIV.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Auth/KerberosIV.thy Fri Oct 05 21:52:39 2001 +0200
@@ -65,10 +65,10 @@
RespLife :: nat
rules
- AuthLife_LB "#2 <= AuthLife"
- ServLife_LB "#2 <= ServLife"
- AutcLife_LB "1' <= AutcLife"
- RespLife_LB "1' <= RespLife"
+ AuthLife_LB "# 2 <= AuthLife"
+ ServLife_LB "# 2 <= ServLife"
+ AutcLife_LB "Suc 0 <= AutcLife"
+ RespLife_LB "Suc 0 <= RespLife"
translations
"CT" == "length"
--- a/src/HOL/Auth/Kerberos_BAN.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Auth/Kerberos_BAN.thy Fri Oct 05 21:52:39 2001 +0200
@@ -30,10 +30,10 @@
rules
(*The ticket should remain fresh for two journeys on the network at least*)
- SesKeyLife_LB "#2 <= SesKeyLife"
+ SesKeyLife_LB "# 2 <= SesKeyLife"
(*The authenticator only for one journey*)
- AutLife_LB "1' <= AutLife"
+ AutLife_LB "Suc 0 <= AutLife"
translations
"CT" == "length"
--- a/src/HOL/Datatype_Universe.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Datatype_Universe.ML Fri Oct 05 21:52:39 2001 +0200
@@ -80,7 +80,7 @@
(** Scons vs Atom **)
-Goalw [Atom_def,Scons_def,Push_Node_def,One_def]
+Goalw [Atom_def,Scons_def,Push_Node_def,One_nat_def]
"Scons M N ~= Atom(a)";
by (rtac notI 1);
by (etac (equalityD2 RS subsetD RS UnE) 1);
@@ -141,11 +141,11 @@
(** Injectiveness of Scons **)
-Goalw [Scons_def,One_def] "Scons M N <= Scons M' N' ==> M<=M'";
+Goalw [Scons_def,One_nat_def] "Scons M N <= Scons M' N' ==> M<=M'";
by (blast_tac (claset() addSDs [Push_Node_inject]) 1);
qed "Scons_inject_lemma1";
-Goalw [Scons_def,One_def] "Scons M N <= Scons M' N' ==> N<=N'";
+Goalw [Scons_def,One_nat_def] "Scons M N <= Scons M' N' ==> N<=N'";
by (blast_tac (claset() addSDs [Push_Node_inject]) 1);
qed "Scons_inject_lemma2";
@@ -252,7 +252,7 @@
by (rtac ntrunc_Atom 1);
qed "ntrunc_Numb";
-Goalw [Scons_def,ntrunc_def,One_def]
+Goalw [Scons_def,ntrunc_def,One_nat_def]
"ntrunc (Suc k) (Scons M N) = Scons (ntrunc k M) (ntrunc k N)";
by (safe_tac (claset() addSIs [imageI]));
by (REPEAT (stac ndepth_Push_Node 3 THEN etac Suc_mono 3));
@@ -266,7 +266,7 @@
(** Injection nodes **)
-Goalw [In0_def] "ntrunc 1' (In0 M) = {}";
+Goalw [In0_def] "ntrunc (Suc 0) (In0 M) = {}";
by (Simp_tac 1);
by (rewtac Scons_def);
by (Blast_tac 1);
@@ -277,7 +277,7 @@
by (Simp_tac 1);
qed "ntrunc_In0";
-Goalw [In1_def] "ntrunc 1' (In1 M) = {}";
+Goalw [In1_def] "ntrunc (Suc 0) (In1 M) = {}";
by (Simp_tac 1);
by (rewtac Scons_def);
by (Blast_tac 1);
@@ -339,7 +339,7 @@
(** Injection **)
-Goalw [In0_def,In1_def,One_def] "In0(M) ~= In1(N)";
+Goalw [In0_def,In1_def,One_nat_def] "In0(M) ~= In1(N)";
by (rtac notI 1);
by (etac (Scons_inject1 RS Numb_inject RS Zero_neq_Suc) 1);
qed "In0_not_In1";
--- a/src/HOL/Datatype_Universe.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Datatype_Universe.thy Fri Oct 05 21:52:39 2001 +0200
@@ -63,7 +63,7 @@
(*S-expression constructors*)
Atom_def "Atom == (%x. {Abs_Node((%k. Inr 0, x))})"
- Scons_def "Scons M N == (Push_Node (Inr 1) ` M) Un (Push_Node (Inr 2) ` N)"
+ Scons_def "Scons M N == (Push_Node (Inr 1) ` M) Un (Push_Node (Inr (Suc 1)) ` N)"
(*Leaf nodes, with arbitrary or nat labels*)
Leaf_def "Leaf == Atom o Inl"
--- a/src/HOL/Divides.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Divides.ML Fri Oct 05 21:52:39 2001 +0200
@@ -65,7 +65,7 @@
by (asm_simp_tac (simpset() addsimps [mod_geq]) 1);
qed "mod_if";
-Goal "m mod 1' = 0";
+Goal "m mod Suc 0 = 0";
by (induct_tac "m" 1);
by (ALLGOALS (asm_simp_tac (simpset() addsimps [mod_geq])));
qed "mod_1";
@@ -387,7 +387,7 @@
(*** Further facts about div and mod ***)
-Goal "m div 1' = m";
+Goal "m div Suc 0 = m";
by (induct_tac "m" 1);
by (ALLGOALS (asm_simp_tac (simpset() addsimps [div_geq])));
qed "div_1";
@@ -529,12 +529,12 @@
qed "dvd_0_left_iff";
AddIffs [dvd_0_left_iff];
-Goalw [dvd_def] "1' dvd k";
+Goalw [dvd_def] "Suc 0 dvd k";
by (Simp_tac 1);
qed "dvd_1_left";
AddIffs [dvd_1_left];
-Goal "(m dvd 1') = (m = 1')";
+Goal "(m dvd Suc 0) = (m = Suc 0)";
by (simp_tac (simpset() addsimps [dvd_def]) 1);
qed "dvd_1_iff_1";
Addsimps [dvd_1_iff_1];
@@ -615,14 +615,14 @@
by (asm_full_simp_tac (simpset() addsimps mult_ac) 1);
qed "dvd_mult_cancel";
-Goal "0<m ==> (m*n dvd m) = (n=1)";
+Goal "0<m ==> (m*n dvd m) = (n = (1::nat))";
by Auto_tac;
by (subgoal_tac "m*n dvd m*1" 1);
by (dtac dvd_mult_cancel 1);
by Auto_tac;
qed "dvd_mult_cancel1";
-Goal "0<m ==> (n*m dvd m) = (n=1)";
+Goal "0<m ==> (n*m dvd m) = (n = (1::nat))";
by (stac mult_commute 1);
by (etac dvd_mult_cancel1 1);
qed "dvd_mult_cancel2";
--- a/src/HOL/Finite.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Finite.ML Fri Oct 05 21:52:39 2001 +0200
@@ -490,7 +490,7 @@
(*** Cardinality of the Powerset ***)
-Goal "finite A ==> card (Pow A) = 2 ^ card A";
+Goal "finite A ==> card (Pow A) = Suc (Suc 0) ^ card A"; (* FIXME numeral 2 (!?) *)
by (etac finite_induct 1);
by (ALLGOALS (asm_simp_tac (simpset() addsimps [Pow_insert])));
by (stac card_Un_disjoint 1);
--- a/src/HOL/GroupTheory/Exponent.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/GroupTheory/Exponent.ML Fri Oct 05 21:52:39 2001 +0200
@@ -8,11 +8,11 @@
val prime_def = thm "prime_def";
-Goalw [prime_def] "p\\<in>prime ==> 1' < p";
+Goalw [prime_def] "p\\<in>prime ==> Suc 0 < p";
by (force_tac (claset(), simpset() addsimps []) 1);
qed "prime_imp_one_less";
-Goal "(p\\<in>prime) = (1'<p & (\\<forall>a b. p dvd a*b --> (p dvd a) | (p dvd b)))";
+Goal "(p\\<in>prime) = (Suc 0 < p & (\\<forall>a b. p dvd a*b --> (p dvd a) | (p dvd b)))";
by (auto_tac (claset(), simpset() addsimps [prime_imp_one_less]));
by (blast_tac (claset() addSDs [thm "prime_dvd_mult"]) 1);
by (auto_tac (claset(), simpset() addsimps [prime_def]));
@@ -201,24 +201,24 @@
qed "div_combine";
(*Lemma for power_dvd_bound*)
-Goal "1' < p ==> Suc n <= p^n";
+Goal "Suc 0 < p ==> Suc n <= p^n";
by (induct_tac "n" 1);
by (Asm_simp_tac 1);
by (Asm_full_simp_tac 1);
-by (subgoal_tac "2*n + #2 <= p * p^n" 1);
+by (subgoal_tac "# 2 * n + # 2 <= p * p^n" 1);
by (Asm_full_simp_tac 1);
-by (subgoal_tac "#2 * p^n <= p * p^n" 1);
+by (subgoal_tac "# 2 * p^n <= p * p^n" 1);
(*?arith_tac should handle all of this!*)
by (rtac order_trans 1);
by (assume_tac 2);
-by (dres_inst_tac [("k","#2")] mult_le_mono2 1);
+by (dres_inst_tac [("k","# 2")] mult_le_mono2 1);
by (Asm_full_simp_tac 1);
by (rtac mult_le_mono1 1);
by (Asm_full_simp_tac 1);
qed "Suc_le_power";
(*An upper bound for the n such that p^n dvd a: needed for GREATEST to exist*)
-Goal "[|p^n dvd a; 1' < p; 0 < a|] ==> n < a";
+Goal "[|p^n dvd a; Suc 0 < p; 0 < a|] ==> n < a";
by (dtac dvd_imp_le 1);
by (dres_inst_tac [("n","n")] Suc_le_power 2);
by Auto_tac;
@@ -267,7 +267,7 @@
Addsimps [exponent_eq_0];
-(* exponent_mult_add, easy inclusion. Could weaken p\\<in>prime to 1'<p *)
+(* exponent_mult_add, easy inclusion. Could weaken p\\<in>prime to Suc 0 < p *)
Goal "[| 0 < a; 0 < b |] \
\ ==> (exponent p a) + (exponent p b) <= exponent p (a * b)";
by (case_tac "p \\<in> prime" 1);
@@ -312,7 +312,7 @@
by (auto_tac (claset() addDs [dvd_mult_left], simpset()));
qed "not_divides_exponent_0";
-Goal "exponent p 1' = 0";
+Goal "exponent p (Suc 0) = 0";
by (case_tac "p \\<in> prime" 1);
by (auto_tac (claset(),
simpset() addsimps [prime_iff, not_divides_exponent_0]));
@@ -357,7 +357,7 @@
Goal "[| 0 < (k::nat); k < p^a; 0 < p; (p^r) dvd (p^a) - k |] ==> r <= a";
-by (res_inst_tac [("m","1'")] p_fac_forw_lemma 1);
+by (res_inst_tac [("m","Suc 0")] p_fac_forw_lemma 1);
by Auto_tac;
qed "r_le_a_forw";
@@ -422,7 +422,7 @@
qed "p_not_div_choose";
-Goal "0 < m ==> exponent p ((p^a * m - 1') choose (p^a - 1')) = 0";
+Goal "0 < m ==> exponent p ((p^a * m - Suc 0) choose (p^a - Suc 0)) = 0";
by (case_tac "p \\<in> prime" 1);
by (Asm_simp_tac 2);
by (forw_inst_tac [("a","a")] zero_less_prime_power 1);
--- a/src/HOL/Hoare/Arith2.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Hoare/Arith2.ML Fri Oct 05 21:52:39 2001 +0200
@@ -63,7 +63,7 @@
(*** pow ***)
-Goal "m mod #2 = 0 ==> ((n::nat)*n)^(m div #2) = n^m";
+Goal "m mod # 2 = 0 ==> ((n::nat)*n)^(m div # 2) = n^m";
by (asm_simp_tac (simpset() addsimps [power_two RS sym, power_mult RS sym,
mult_div_cancel]) 1);
qed "sq_pow_div2";
--- a/src/HOL/Hoare/Examples.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Hoare/Examples.ML Fri Oct 05 21:52:39 2001 +0200
@@ -13,7 +13,7 @@
\ m := 0; s := 0; \
\ WHILE m~=a \
\ INV {s=m*b & a=A & b=B} \
-\ DO s := s+b; m := m+1 OD \
+\ DO s := s+b; m := m+(1::nat) OD \
\ {s = A*B}";
by (hoare_tac (Asm_full_simp_tac) 1);
qed "multiply_by_add";
@@ -50,9 +50,9 @@
Goal "|- VARS a b x y. \
\ {0<A & 0<B & a=A & b=B & x=B & y=A} \
\ WHILE a ~= b \
-\ INV {0<a & 0<b & gcd A B = gcd a b & #2*A*B = a*x + b*y} \
+\ INV {0<a & 0<b & gcd A B = gcd a b & # 2*A*B = a*x + b*y} \
\ DO IF a<b THEN (b := b-a; x := x+y) ELSE (a := a-b; y := y+x) FI OD \
-\ {a = gcd A B & #2*A*B = a*(x+y)}";
+\ {a = gcd A B & # 2*A*B = a*(x+y)}";
by (hoare_tac (K all_tac) 1);
by(Asm_simp_tac 1);
by(asm_simp_tac (simpset() addsimps
@@ -65,13 +65,13 @@
Goal "|- VARS a b c. \
\ {a=A & b=B} \
-\ c := 1; \
+\ c := (1::nat); \
\ WHILE b ~= 0 \
\ INV {A^B = c * a^b} \
-\ DO WHILE b mod #2 = 0 \
+\ DO WHILE b mod # 2 = 0 \
\ INV {A^B = c * a^b} \
-\ DO a := a*a; b := b div #2 OD; \
-\ c := c*a; b := b-1 \
+\ DO a := a*a; b := b div # 2 OD; \
+\ c := c*a; b := b - 1 \
\ OD \
\ {c = A^B}";
by (hoare_tac (Asm_full_simp_tac) 1);
@@ -87,7 +87,7 @@
\ b := 1; \
\ WHILE a ~= 0 \
\ INV {fac A = b * fac a} \
-\ DO b := b*a; a := a-1 OD \
+\ DO b := b*a; a := a - 1 OD \
\ {b = fac A}";
by (hoare_tac (asm_full_simp_tac (simpset() addsplits [nat_diff_split])) 1);
by Auto_tac;
@@ -99,7 +99,7 @@
Goal "|- VARS r x. \
\ {True} \
-\ x := X; r := 0; \
+\ x := X; r := (0::nat); \
\ WHILE (r+1)*(r+1) <= x \
\ INV {r*r <= x & x=X} \
\ DO r := r+1 OD \
@@ -111,10 +111,10 @@
Goal "|- VARS u w r x. \
\ {True} \
-\ x := X; u := 1; w := 1; r := 0; \
+\ x := X; u := 1; w := 1; r := (0::nat); \
\ WHILE w <= x \
\ INV {u = r+r+1 & w = (r+1)*(r+1) & r*r <= x & x=X} \
-\ DO r := r+1; w := w+u+2; u := u+2 OD \
+\ DO r := r + 1; w := w + u + # 2; u := u + # 2 OD \
\ {r*r <= X & X < (r+1)*(r+1)}";
by (hoare_tac (SELECT_GOAL Auto_tac) 1);
qed "sqrt_without_multiplication";
@@ -175,7 +175,7 @@
Ambiguity warnings of parser are due to := being used
both for assignment and list update.
*)
-Goal "m - 1' < n ==> m < Suc n";
+Goal "m - Suc 0 < n ==> m < Suc n";
by (arith_tac 1);
qed "lemma";
@@ -184,7 +184,7 @@
\ geq == %A i. !k. i<k & k<length A --> pivot <= A!k |] ==> \
\ |- VARS A u l.\
\ {0 < length(A::('a::order)list)} \
-\ l := 0; u := length A - 1'; \
+\ l := 0; u := length A - Suc 0; \
\ WHILE l <= u \
\ INV {leq A l & geq A u & u<length A & l<=length A} \
\ DO WHILE l < length A & A!l <= pivot \
@@ -192,7 +192,7 @@
\ DO l := l+1 OD; \
\ WHILE 0 < u & pivot <= A!u \
\ INV {leq A l & geq A u & u<length A & l<=length A} \
-\ DO u := u-1 OD; \
+\ DO u := u - 1 OD; \
\ IF l <= u THEN A := A[l := A!u, u := A!l] ELSE SKIP FI \
\ OD \
\ {leq A u & (!k. u<k & k<l --> A!k = pivot) & geq A l}";
--- a/src/HOL/Hyperreal/HRealAbs.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Hyperreal/HRealAbs.ML Fri Oct 05 21:52:39 2001 +0200
@@ -32,28 +32,28 @@
(adapted version of previously proved theorems about abs)
------------------------------------------------------------*)
-Goal "abs (#0::hypreal) = #0";
+Goal "abs (Numeral0::hypreal) = Numeral0";
by (simp_tac (simpset() addsimps [hrabs_def]) 1);
qed "hrabs_zero";
Addsimps [hrabs_zero];
-Goal "(#0::hypreal)<=x ==> abs x = x";
+Goal "(Numeral0::hypreal)<=x ==> abs x = x";
by (asm_simp_tac (simpset() addsimps [hrabs_def]) 1);
qed "hrabs_eqI1";
-Goal "(#0::hypreal)<x ==> abs x = x";
+Goal "(Numeral0::hypreal)<x ==> abs x = x";
by (asm_simp_tac (simpset() addsimps [order_less_imp_le, hrabs_eqI1]) 1);
qed "hrabs_eqI2";
-Goal "x<(#0::hypreal) ==> abs x = -x";
+Goal "x<(Numeral0::hypreal) ==> abs x = -x";
by (asm_simp_tac (simpset() addsimps [hypreal_le_def, hrabs_def]) 1);
qed "hrabs_minus_eqI2";
-Goal "x<=(#0::hypreal) ==> abs x = -x";
+Goal "x<=(Numeral0::hypreal) ==> abs x = -x";
by (auto_tac (claset() addDs [order_antisym], simpset() addsimps [hrabs_def]));
qed "hrabs_minus_eqI1";
-Goal "(#0::hypreal)<= abs x";
+Goal "(Numeral0::hypreal)<= abs x";
by (auto_tac (claset() addDs [hypreal_minus_zero_less_iff RS iffD2,
hypreal_less_asym],
simpset() addsimps [hypreal_le_def, hrabs_def]));
@@ -66,7 +66,7 @@
qed "hrabs_idempotent";
Addsimps [hrabs_idempotent];
-Goalw [hrabs_def] "(abs x = (#0::hypreal)) = (x=#0)";
+Goalw [hrabs_def] "(abs x = (Numeral0::hypreal)) = (x=Numeral0)";
by (Simp_tac 1);
qed "hrabs_zero_iff";
AddIffs [hrabs_zero_iff];
@@ -90,7 +90,7 @@
Addsimps [hrabs_mult];
Goal "abs(inverse(x)) = inverse(abs(x::hypreal))";
-by (hypreal_div_undefined_case_tac "x=#0" 1);
+by (hypreal_div_undefined_case_tac "x=Numeral0" 1);
by (simp_tac (simpset() addsimps [HYPREAL_DIVIDE_ZERO]) 1);
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
by (auto_tac (claset(),
@@ -128,10 +128,10 @@
qed "hrabs_add_less";
Goal "[| abs x<r; abs y<s |] ==> abs x * abs y < r * (s::hypreal)";
-by (subgoal_tac "#0 < r" 1);
+by (subgoal_tac "Numeral0 < r" 1);
by (asm_full_simp_tac (simpset() addsimps [hrabs_def]
addsplits [split_if_asm]) 2);
-by (case_tac "y = #0" 1);
+by (case_tac "y = Numeral0" 1);
by (asm_full_simp_tac (simpset() addsimps [hypreal_0_less_mult_iff]) 1);
by (rtac hypreal_mult_less_mono 1);
by (auto_tac (claset(),
@@ -139,18 +139,18 @@
addsplits [split_if_asm]));
qed "hrabs_mult_less";
-Goal "((#0::hypreal) < abs x) = (x ~= 0)";
+Goal "((Numeral0::hypreal) < abs x) = (x ~= 0)";
by (simp_tac (simpset() addsimps [hrabs_def]) 1);
by (arith_tac 1);
qed "hypreal_0_less_abs_iff";
Addsimps [hypreal_0_less_abs_iff];
-Goal "abs x < r ==> (#0::hypreal) < r";
+Goal "abs x < r ==> (Numeral0::hypreal) < r";
by (blast_tac (claset() addSIs [order_le_less_trans, hrabs_ge_zero]) 1);
qed "hrabs_less_gt_zero";
Goal "abs x = (x::hypreal) | abs x = -x";
-by (cut_inst_tac [("x","#0"),("y","x")] hypreal_linear 1);
+by (cut_inst_tac [("x","Numeral0"),("y","x")] hypreal_linear 1);
by (fast_tac (claset() addIs [hrabs_eqI2,hrabs_minus_eqI2,
hrabs_zero]) 1);
qed "hrabs_disj";
@@ -247,13 +247,13 @@
(*"neg" is used in rewrite rules for binary comparisons*)
Goal "hypreal_of_nat (number_of v :: nat) = \
-\ (if neg (number_of v) then #0 \
+\ (if neg (number_of v) then Numeral0 \
\ else (number_of v :: hypreal))";
by (simp_tac (simpset() addsimps [hypreal_of_nat_def]) 1);
qed "hypreal_of_nat_number_of";
Addsimps [hypreal_of_nat_number_of];
-Goal "hypreal_of_nat 0 = #0";
+Goal "hypreal_of_nat 0 = Numeral0";
by (simp_tac (simpset() delsimps [numeral_0_eq_0]
addsimps [numeral_0_eq_0 RS sym]) 1);
qed "hypreal_of_nat_zero";
--- a/src/HOL/Hyperreal/HSeries.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Hyperreal/HSeries.ML Fri Oct 05 21:52:39 2001 +0200
@@ -35,7 +35,7 @@
(* Theorem corresponding to base case in def of sumr *)
Goalw [hypnat_zero_def]
- "sumhr (m,0,f) = #0";
+ "sumhr (m,0,f) = Numeral0";
by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
by (auto_tac (claset(),
simpset() addsimps [sumhr, symmetric hypreal_zero_def]));
@@ -44,7 +44,7 @@
(* Theorem corresponding to recursive case in def of sumr *)
Goalw [hypnat_one_def]
- "sumhr(m,n+1hn,f) = (if n + 1hn <= m then #0 \
+ "sumhr(m,n+1hn,f) = (if n + 1hn <= m then Numeral0 \
\ else sumhr(m,n,f) + (*fNat* f) n)";
by (simp_tac (HOL_ss addsimps
[zero_eq_numeral_0 RS sym, hypreal_zero_def]) 1);
@@ -55,7 +55,7 @@
by (ALLGOALS(Ultra_tac));
qed "sumhr_if";
-Goalw [hypnat_one_def] "sumhr (n + 1hn, n, f) = #0";
+Goalw [hypnat_one_def] "sumhr (n + 1hn, n, f) = Numeral0";
by (simp_tac (HOL_ss addsimps
[zero_eq_numeral_0 RS sym, hypreal_zero_def]) 1);
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
@@ -64,7 +64,7 @@
qed "sumhr_Suc_zero";
Addsimps [sumhr_Suc_zero];
-Goal "sumhr (n,n,f) = #0";
+Goal "sumhr (n,n,f) = Numeral0";
by (simp_tac (HOL_ss addsimps
[zero_eq_numeral_0 RS sym, hypreal_zero_def]) 1);
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
@@ -80,7 +80,7 @@
qed "sumhr_Suc";
Addsimps [sumhr_Suc];
-Goal "sumhr(m+k,k,f) = #0";
+Goal "sumhr(m+k,k,f) = Numeral0";
by (simp_tac (HOL_ss addsimps
[zero_eq_numeral_0 RS sym, hypreal_zero_def]) 1);
by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
@@ -156,7 +156,7 @@
hypreal_minus,sumr_add RS sym]) 1);
qed "sumhr_add_mult_const";
-Goal "n < m ==> sumhr (m,n,f) = #0";
+Goal "n < m ==> sumhr (m,n,f) = Numeral0";
by (simp_tac (HOL_ss addsimps
[zero_eq_numeral_0 RS sym, hypreal_zero_def]) 1);
by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
@@ -185,13 +185,13 @@
by summing to some infinite hypernatural (such as whn)
-----------------------------------------------------------------*)
Goalw [hypnat_omega_def,hypnat_zero_def]
- "sumhr(0,whn,%i. #1) = hypreal_of_hypnat whn";
+ "sumhr(0,whn,%i. Numeral1) = hypreal_of_hypnat whn";
by (auto_tac (claset(),
simpset() addsimps [sumhr, hypreal_of_hypnat]));
qed "sumhr_hypreal_of_hypnat_omega";
Goalw [hypnat_omega_def,hypnat_zero_def,omega_def]
- "sumhr(0, whn, %i. #1) = omega - #1";
+ "sumhr(0, whn, %i. Numeral1) = omega - Numeral1";
by (simp_tac (HOL_ss addsimps
[one_eq_numeral_1 RS sym, hypreal_one_def]) 1);
by (auto_tac (claset(),
@@ -199,7 +199,7 @@
qed "sumhr_hypreal_omega_minus_one";
Goalw [hypnat_zero_def, hypnat_omega_def]
- "sumhr(0, whn + whn, %i. (-#1) ^ (i+1)) = #0";
+ "sumhr(0, whn + whn, %i. (-Numeral1) ^ (i+1)) = Numeral0";
by (simp_tac (HOL_ss addsimps
[zero_eq_numeral_0 RS sym, hypreal_zero_def]) 1);
by (simp_tac (simpset() addsimps [sumhr,hypnat_add,double_lemma]
@@ -223,7 +223,7 @@
qed "starfunNat_sumr";
Goal "sumhr (0, M, f) @= sumhr (0, N, f) \
-\ ==> abs (sumhr (M, N, f)) @= #0";
+\ ==> abs (sumhr (M, N, f)) @= Numeral0";
by (cut_inst_tac [("x","M"),("y","N")] hypnat_linear 1);
by (auto_tac (claset(), simpset() addsimps [approx_refl]));
by (dtac (approx_sym RS (approx_minus_iff RS iffD1)) 1);
@@ -265,12 +265,12 @@
sums_unique]) 1);
qed "NSsums_unique";
-Goal "ALL m. n <= Suc m --> f(m) = #0 ==> f NSsums (sumr 0 n f)";
+Goal "ALL m. n <= Suc m --> f(m) = Numeral0 ==> f NSsums (sumr 0 n f)";
by (asm_simp_tac (simpset() addsimps [sums_NSsums_iff RS sym, series_zero]) 1);
qed "NSseries_zero";
Goal "NSsummable f = \
-\ (ALL M: HNatInfinite. ALL N: HNatInfinite. abs (sumhr(M,N,f)) @= #0)";
+\ (ALL M: HNatInfinite. ALL N: HNatInfinite. abs (sumhr(M,N,f)) @= Numeral0)";
by (auto_tac (claset(),
simpset() addsimps [summable_NSsummable_iff RS sym,
summable_convergent_sumr_iff, convergent_NSconvergent_iff,
@@ -287,7 +287,7 @@
(*-------------------------------------------------------------------
Terms of a convergent series tend to zero
-------------------------------------------------------------------*)
-Goalw [NSLIMSEQ_def] "NSsummable f ==> f ----NS> #0";
+Goalw [NSLIMSEQ_def] "NSsummable f ==> f ----NS> Numeral0";
by (auto_tac (claset(), simpset() addsimps [NSsummable_NSCauchy]));
by (dtac bspec 1 THEN Auto_tac);
by (dres_inst_tac [("x","N + 1hn")] bspec 1);
@@ -297,7 +297,7 @@
qed "NSsummable_NSLIMSEQ_zero";
(* Easy to prove stsandard case now *)
-Goal "summable f ==> f ----> #0";
+Goal "summable f ==> f ----> Numeral0";
by (auto_tac (claset(),
simpset() addsimps [summable_NSsummable_iff,
LIMSEQ_NSLIMSEQ_iff, NSsummable_NSLIMSEQ_zero]));
--- a/src/HOL/Hyperreal/HyperArith0.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Hyperreal/HyperArith0.ML Fri Oct 05 21:52:39 2001 +0200
@@ -8,7 +8,7 @@
Also, common factor cancellation
*)
-Goal "((x * y = #0) = (x = #0 | y = (#0::hypreal)))";
+Goal "((x * y = Numeral0) = (x = Numeral0 | y = (Numeral0::hypreal)))";
by Auto_tac;
by (cut_inst_tac [("x","x"),("y","y")] hypreal_mult_zero_disj 1);
by Auto_tac;
@@ -17,13 +17,13 @@
(** Division and inverse **)
-Goal "#0/x = (#0::hypreal)";
+Goal "Numeral0/x = (Numeral0::hypreal)";
by (simp_tac (simpset() addsimps [hypreal_divide_def]) 1);
qed "hypreal_0_divide";
Addsimps [hypreal_0_divide];
-Goal "((#0::hypreal) < inverse x) = (#0 < x)";
-by (case_tac "x=#0" 1);
+Goal "((Numeral0::hypreal) < inverse x) = (Numeral0 < x)";
+by (case_tac "x=Numeral0" 1);
by (asm_simp_tac (HOL_ss addsimps [rename_numerals HYPREAL_INVERSE_ZERO]) 1);
by (auto_tac (claset() addDs [hypreal_inverse_less_0],
simpset() addsimps [linorder_neq_iff,
@@ -31,8 +31,8 @@
qed "hypreal_0_less_inverse_iff";
Addsimps [hypreal_0_less_inverse_iff];
-Goal "(inverse x < (#0::hypreal)) = (x < #0)";
-by (case_tac "x=#0" 1);
+Goal "(inverse x < (Numeral0::hypreal)) = (x < Numeral0)";
+by (case_tac "x=Numeral0" 1);
by (asm_simp_tac (HOL_ss addsimps [rename_numerals HYPREAL_INVERSE_ZERO]) 1);
by (auto_tac (claset() addDs [hypreal_inverse_less_0],
simpset() addsimps [linorder_neq_iff,
@@ -40,49 +40,49 @@
qed "hypreal_inverse_less_0_iff";
Addsimps [hypreal_inverse_less_0_iff];
-Goal "((#0::hypreal) <= inverse x) = (#0 <= x)";
+Goal "((Numeral0::hypreal) <= inverse x) = (Numeral0 <= x)";
by (simp_tac (simpset() addsimps [linorder_not_less RS sym]) 1);
qed "hypreal_0_le_inverse_iff";
Addsimps [hypreal_0_le_inverse_iff];
-Goal "(inverse x <= (#0::hypreal)) = (x <= #0)";
+Goal "(inverse x <= (Numeral0::hypreal)) = (x <= Numeral0)";
by (simp_tac (simpset() addsimps [linorder_not_less RS sym]) 1);
qed "hypreal_inverse_le_0_iff";
Addsimps [hypreal_inverse_le_0_iff];
-Goalw [hypreal_divide_def] "x/(#0::hypreal) = #0";
+Goalw [hypreal_divide_def] "x/(Numeral0::hypreal) = Numeral0";
by (stac (rename_numerals HYPREAL_INVERSE_ZERO) 1);
by (Simp_tac 1);
qed "HYPREAL_DIVIDE_ZERO";
-Goal "inverse (x::hypreal) = #1/x";
+Goal "inverse (x::hypreal) = Numeral1/x";
by (simp_tac (simpset() addsimps [hypreal_divide_def]) 1);
qed "hypreal_inverse_eq_divide";
-Goal "((#0::hypreal) < x/y) = (#0 < x & #0 < y | x < #0 & y < #0)";
+Goal "((Numeral0::hypreal) < x/y) = (Numeral0 < x & Numeral0 < y | x < Numeral0 & y < Numeral0)";
by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_0_less_mult_iff]) 1);
qed "hypreal_0_less_divide_iff";
Addsimps [inst "x" "number_of ?w" hypreal_0_less_divide_iff];
-Goal "(x/y < (#0::hypreal)) = (#0 < x & y < #0 | x < #0 & #0 < y)";
+Goal "(x/y < (Numeral0::hypreal)) = (Numeral0 < x & y < Numeral0 | x < Numeral0 & Numeral0 < y)";
by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_less_0_iff]) 1);
qed "hypreal_divide_less_0_iff";
Addsimps [inst "x" "number_of ?w" hypreal_divide_less_0_iff];
-Goal "((#0::hypreal) <= x/y) = ((x <= #0 | #0 <= y) & (#0 <= x | y <= #0))";
+Goal "((Numeral0::hypreal) <= x/y) = ((x <= Numeral0 | Numeral0 <= y) & (Numeral0 <= x | y <= Numeral0))";
by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_0_le_mult_iff]) 1);
by Auto_tac;
qed "hypreal_0_le_divide_iff";
Addsimps [inst "x" "number_of ?w" hypreal_0_le_divide_iff];
-Goal "(x/y <= (#0::hypreal)) = ((x <= #0 | y <= #0) & (#0 <= x | #0 <= y))";
+Goal "(x/y <= (Numeral0::hypreal)) = ((x <= Numeral0 | y <= Numeral0) & (Numeral0 <= x | Numeral0 <= y))";
by (simp_tac (simpset() addsimps [hypreal_divide_def,
hypreal_mult_le_0_iff]) 1);
by Auto_tac;
qed "hypreal_divide_le_0_iff";
Addsimps [inst "x" "number_of ?w" hypreal_divide_le_0_iff];
-Goal "(inverse(x::hypreal) = #0) = (x = #0)";
+Goal "(inverse(x::hypreal) = Numeral0) = (x = Numeral0)";
by (auto_tac (claset(),
simpset() addsimps [rename_numerals HYPREAL_INVERSE_ZERO]));
by (rtac ccontr 1);
@@ -90,12 +90,12 @@
qed "hypreal_inverse_zero_iff";
Addsimps [hypreal_inverse_zero_iff];
-Goal "(x/y = #0) = (x=#0 | y=(#0::hypreal))";
+Goal "(x/y = Numeral0) = (x=Numeral0 | y=(Numeral0::hypreal))";
by (auto_tac (claset(), simpset() addsimps [hypreal_divide_def]));
qed "hypreal_divide_eq_0_iff";
Addsimps [hypreal_divide_eq_0_iff];
-Goal "h ~= (#0::hypreal) ==> h/h = #1";
+Goal "h ~= (Numeral0::hypreal) ==> h/h = Numeral1";
by (asm_simp_tac
(simpset() addsimps [hypreal_divide_def, hypreal_mult_inverse_left]) 1);
qed "hypreal_divide_self_eq";
@@ -140,7 +140,7 @@
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [hypreal_mult_commute])));
qed "hypreal_mult_le_mono2_neg";
-Goal "(m*k < n*k) = (((#0::hypreal) < k & m<n) | (k < #0 & n<m))";
+Goal "(m*k < n*k) = (((Numeral0::hypreal) < k & m<n) | (k < Numeral0 & n<m))";
by (case_tac "k = (0::hypreal)" 1);
by (auto_tac (claset(),
simpset() addsimps [linorder_neq_iff,
@@ -155,32 +155,32 @@
hypreal_mult_le_mono1_neg]));
qed "hypreal_mult_less_cancel2";
-Goal "(m*k <= n*k) = (((#0::hypreal) < k --> m<=n) & (k < #0 --> n<=m))";
+Goal "(m*k <= n*k) = (((Numeral0::hypreal) < k --> m<=n) & (k < Numeral0 --> n<=m))";
by (simp_tac (simpset() addsimps [linorder_not_less RS sym,
hypreal_mult_less_cancel2]) 1);
qed "hypreal_mult_le_cancel2";
-Goal "(k*m < k*n) = (((#0::hypreal) < k & m<n) | (k < #0 & n<m))";
+Goal "(k*m < k*n) = (((Numeral0::hypreal) < k & m<n) | (k < Numeral0 & n<m))";
by (simp_tac (simpset() addsimps [inst "z" "k" hypreal_mult_commute,
hypreal_mult_less_cancel2]) 1);
qed "hypreal_mult_less_cancel1";
-Goal "!!k::hypreal. (k*m <= k*n) = ((#0 < k --> m<=n) & (k < #0 --> n<=m))";
+Goal "!!k::hypreal. (k*m <= k*n) = ((Numeral0 < k --> m<=n) & (k < Numeral0 --> n<=m))";
by (simp_tac (simpset() addsimps [linorder_not_less RS sym,
hypreal_mult_less_cancel1]) 1);
qed "hypreal_mult_le_cancel1";
-Goal "!!k::hypreal. (k*m = k*n) = (k = #0 | m=n)";
+Goal "!!k::hypreal. (k*m = k*n) = (k = Numeral0 | m=n)";
by (case_tac "k=0" 1);
by (auto_tac (claset(), simpset() addsimps [hypreal_mult_left_cancel]));
qed "hypreal_mult_eq_cancel1";
-Goal "!!k::hypreal. (m*k = n*k) = (k = #0 | m=n)";
+Goal "!!k::hypreal. (m*k = n*k) = (k = Numeral0 | m=n)";
by (case_tac "k=0" 1);
by (auto_tac (claset(), simpset() addsimps [hypreal_mult_right_cancel]));
qed "hypreal_mult_eq_cancel2";
-Goal "!!k::hypreal. k~=#0 ==> (k*m) / (k*n) = (m/n)";
+Goal "!!k::hypreal. k~=Numeral0 ==> (k*m) / (k*n) = (m/n)";
by (asm_simp_tac
(simpset() addsimps [hypreal_divide_def, hypreal_inverse_distrib]) 1);
by (subgoal_tac "k * m * (inverse k * inverse n) = \
@@ -190,7 +190,7 @@
qed "hypreal_mult_div_cancel1";
(*For ExtractCommonTerm*)
-Goal "(k*m) / (k*n) = (if k = (#0::hypreal) then #0 else m/n)";
+Goal "(k*m) / (k*n) = (if k = (Numeral0::hypreal) then Numeral0 else m/n)";
by (simp_tac (simpset() addsimps [hypreal_mult_div_cancel1]) 1);
qed "hypreal_mult_div_cancel_disj";
@@ -288,34 +288,34 @@
set trace_simp;
fun test s = (Goal s; by (Simp_tac 1));
-test "#0 <= (y::hypreal) * #-2";
-test "#9*x = #12 * (y::hypreal)";
-test "(#9*x) / (#12 * (y::hypreal)) = z";
-test "#9*x < #12 * (y::hypreal)";
-test "#9*x <= #12 * (y::hypreal)";
+test "Numeral0 <= (y::hypreal) * # -2";
+test "# 9*x = # 12 * (y::hypreal)";
+test "(# 9*x) / (# 12 * (y::hypreal)) = z";
+test "# 9*x < # 12 * (y::hypreal)";
+test "# 9*x <= # 12 * (y::hypreal)";
-test "#-99*x = #132 * (y::hypreal)";
-test "(#-99*x) / (#132 * (y::hypreal)) = z";
-test "#-99*x < #132 * (y::hypreal)";
-test "#-99*x <= #132 * (y::hypreal)";
+test "# -99*x = # 123 * (y::hypreal)";
+test "(# -99*x) / (# 123 * (y::hypreal)) = z";
+test "# -99*x < # 123 * (y::hypreal)";
+test "# -99*x <= # 123 * (y::hypreal)";
-test "#999*x = #-396 * (y::hypreal)";
-test "(#999*x) / (#-396 * (y::hypreal)) = z";
-test "#999*x < #-396 * (y::hypreal)";
-test "#999*x <= #-396 * (y::hypreal)";
+test "# 999*x = # -396 * (y::hypreal)";
+test "(# 999*x) / (# -396 * (y::hypreal)) = z";
+test "# 999*x < # -396 * (y::hypreal)";
+test "# 999*x <= # -396 * (y::hypreal)";
-test "#-99*x = #-81 * (y::hypreal)";
-test "(#-99*x) / (#-81 * (y::hypreal)) = z";
-test "#-99*x <= #-81 * (y::hypreal)";
-test "#-99*x < #-81 * (y::hypreal)";
+test "# -99*x = # -81 * (y::hypreal)";
+test "(# -99*x) / (# -81 * (y::hypreal)) = z";
+test "# -99*x <= # -81 * (y::hypreal)";
+test "# -99*x < # -81 * (y::hypreal)";
-test "#-2 * x = #-1 * (y::hypreal)";
-test "#-2 * x = -(y::hypreal)";
-test "(#-2 * x) / (#-1 * (y::hypreal)) = z";
-test "#-2 * x < -(y::hypreal)";
-test "#-2 * x <= #-1 * (y::hypreal)";
-test "-x < #-23 * (y::hypreal)";
-test "-x <= #-23 * (y::hypreal)";
+test "# -2 * x = # -1 * (y::hypreal)";
+test "# -2 * x = -(y::hypreal)";
+test "(# -2 * x) / (# -1 * (y::hypreal)) = z";
+test "# -2 * x < -(y::hypreal)";
+test "# -2 * x <= # -1 * (y::hypreal)";
+test "-x < # -23 * (y::hypreal)";
+test "-x <= # -23 * (y::hypreal)";
*)
@@ -391,7 +391,7 @@
(*** Simplification of inequalities involving literal divisors ***)
-Goal "#0<z ==> ((x::hypreal) <= y/z) = (x*z <= y)";
+Goal "Numeral0<z ==> ((x::hypreal) <= y/z) = (x*z <= y)";
by (subgoal_tac "(x*z <= y) = (x*z <= (y/z)*z)" 1);
by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2);
by (etac ssubst 1);
@@ -400,7 +400,7 @@
qed "pos_hypreal_le_divide_eq";
Addsimps [inst "z" "number_of ?w" pos_hypreal_le_divide_eq];
-Goal "z<#0 ==> ((x::hypreal) <= y/z) = (y <= x*z)";
+Goal "z<Numeral0 ==> ((x::hypreal) <= y/z) = (y <= x*z)";
by (subgoal_tac "(y <= x*z) = ((y/z)*z <= x*z)" 1);
by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2);
by (etac ssubst 1);
@@ -409,7 +409,7 @@
qed "neg_hypreal_le_divide_eq";
Addsimps [inst "z" "number_of ?w" neg_hypreal_le_divide_eq];
-Goal "#0<z ==> (y/z <= (x::hypreal)) = (y <= x*z)";
+Goal "Numeral0<z ==> (y/z <= (x::hypreal)) = (y <= x*z)";
by (subgoal_tac "(y <= x*z) = ((y/z)*z <= x*z)" 1);
by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2);
by (etac ssubst 1);
@@ -418,7 +418,7 @@
qed "pos_hypreal_divide_le_eq";
Addsimps [inst "z" "number_of ?w" pos_hypreal_divide_le_eq];
-Goal "z<#0 ==> (y/z <= (x::hypreal)) = (x*z <= y)";
+Goal "z<Numeral0 ==> (y/z <= (x::hypreal)) = (x*z <= y)";
by (subgoal_tac "(x*z <= y) = (x*z <= (y/z)*z)" 1);
by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2);
by (etac ssubst 1);
@@ -427,7 +427,7 @@
qed "neg_hypreal_divide_le_eq";
Addsimps [inst "z" "number_of ?w" neg_hypreal_divide_le_eq];
-Goal "#0<z ==> ((x::hypreal) < y/z) = (x*z < y)";
+Goal "Numeral0<z ==> ((x::hypreal) < y/z) = (x*z < y)";
by (subgoal_tac "(x*z < y) = (x*z < (y/z)*z)" 1);
by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2);
by (etac ssubst 1);
@@ -436,7 +436,7 @@
qed "pos_hypreal_less_divide_eq";
Addsimps [inst "z" "number_of ?w" pos_hypreal_less_divide_eq];
-Goal "z<#0 ==> ((x::hypreal) < y/z) = (y < x*z)";
+Goal "z<Numeral0 ==> ((x::hypreal) < y/z) = (y < x*z)";
by (subgoal_tac "(y < x*z) = ((y/z)*z < x*z)" 1);
by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2);
by (etac ssubst 1);
@@ -445,7 +445,7 @@
qed "neg_hypreal_less_divide_eq";
Addsimps [inst "z" "number_of ?w" neg_hypreal_less_divide_eq];
-Goal "#0<z ==> (y/z < (x::hypreal)) = (y < x*z)";
+Goal "Numeral0<z ==> (y/z < (x::hypreal)) = (y < x*z)";
by (subgoal_tac "(y < x*z) = ((y/z)*z < x*z)" 1);
by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2);
by (etac ssubst 1);
@@ -454,7 +454,7 @@
qed "pos_hypreal_divide_less_eq";
Addsimps [inst "z" "number_of ?w" pos_hypreal_divide_less_eq];
-Goal "z<#0 ==> (y/z < (x::hypreal)) = (x*z < y)";
+Goal "z<Numeral0 ==> (y/z < (x::hypreal)) = (x*z < y)";
by (subgoal_tac "(x*z < y) = (x*z < (y/z)*z)" 1);
by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2);
by (etac ssubst 1);
@@ -463,7 +463,7 @@
qed "neg_hypreal_divide_less_eq";
Addsimps [inst "z" "number_of ?w" neg_hypreal_divide_less_eq];
-Goal "z~=#0 ==> ((x::hypreal) = y/z) = (x*z = y)";
+Goal "z~=Numeral0 ==> ((x::hypreal) = y/z) = (x*z = y)";
by (subgoal_tac "(x*z = y) = (x*z = (y/z)*z)" 1);
by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2);
by (etac ssubst 1);
@@ -472,7 +472,7 @@
qed "hypreal_eq_divide_eq";
Addsimps [inst "z" "number_of ?w" hypreal_eq_divide_eq];
-Goal "z~=#0 ==> (y/z = (x::hypreal)) = (y = x*z)";
+Goal "z~=Numeral0 ==> (y/z = (x::hypreal)) = (y = x*z)";
by (subgoal_tac "(y = x*z) = ((y/z)*z = x*z)" 1);
by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2);
by (etac ssubst 1);
@@ -481,21 +481,21 @@
qed "hypreal_divide_eq_eq";
Addsimps [inst "z" "number_of ?w" hypreal_divide_eq_eq];
-Goal "(m/k = n/k) = (k = #0 | m = (n::hypreal))";
-by (case_tac "k=#0" 1);
+Goal "(m/k = n/k) = (k = Numeral0 | m = (n::hypreal))";
+by (case_tac "k=Numeral0" 1);
by (asm_simp_tac (simpset() addsimps [HYPREAL_DIVIDE_ZERO]) 1);
by (asm_simp_tac (simpset() addsimps [hypreal_divide_eq_eq, hypreal_eq_divide_eq,
hypreal_mult_eq_cancel2]) 1);
qed "hypreal_divide_eq_cancel2";
-Goal "(k/m = k/n) = (k = #0 | m = (n::hypreal))";
-by (case_tac "m=#0 | n = #0" 1);
+Goal "(k/m = k/n) = (k = Numeral0 | m = (n::hypreal))";
+by (case_tac "m=Numeral0 | n = Numeral0" 1);
by (auto_tac (claset(),
simpset() addsimps [HYPREAL_DIVIDE_ZERO, hypreal_divide_eq_eq,
hypreal_eq_divide_eq, hypreal_mult_eq_cancel1]));
qed "hypreal_divide_eq_cancel1";
-Goal "[| #0 < r; #0 < x|] ==> (inverse x < inverse (r::hypreal)) = (r < x)";
+Goal "[| Numeral0 < r; Numeral0 < x|] ==> (inverse x < inverse (r::hypreal)) = (r < x)";
by (auto_tac (claset() addIs [hypreal_inverse_less_swap], simpset()));
by (res_inst_tac [("t","r")] (hypreal_inverse_inverse RS subst) 1);
by (res_inst_tac [("t","x")] (hypreal_inverse_inverse RS subst) 1);
@@ -504,30 +504,30 @@
addsimps [hypreal_inverse_gt_zero]));
qed "hypreal_inverse_less_iff";
-Goal "[| #0 < r; #0 < x|] ==> (inverse x <= inverse r) = (r <= (x::hypreal))";
+Goal "[| Numeral0 < r; Numeral0 < x|] ==> (inverse x <= inverse r) = (r <= (x::hypreal))";
by (asm_simp_tac (simpset() addsimps [linorder_not_less RS sym,
hypreal_inverse_less_iff]) 1);
qed "hypreal_inverse_le_iff";
(** Division by 1, -1 **)
-Goal "(x::hypreal)/#1 = x";
+Goal "(x::hypreal)/Numeral1 = x";
by (simp_tac (simpset() addsimps [hypreal_divide_def]) 1);
qed "hypreal_divide_1";
Addsimps [hypreal_divide_1];
-Goal "x/#-1 = -(x::hypreal)";
+Goal "x/# -1 = -(x::hypreal)";
by (Simp_tac 1);
qed "hypreal_divide_minus1";
Addsimps [hypreal_divide_minus1];
-Goal "#-1/(x::hypreal) = - (#1/x)";
+Goal "# -1/(x::hypreal) = - (Numeral1/x)";
by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_minus_inverse]) 1);
qed "hypreal_minus1_divide";
Addsimps [hypreal_minus1_divide];
-Goal "[| (#0::hypreal) < d1; #0 < d2 |] ==> EX e. #0 < e & e < d1 & e < d2";
-by (res_inst_tac [("x","(min d1 d2)/#2")] exI 1);
+Goal "[| (Numeral0::hypreal) < d1; Numeral0 < d2 |] ==> EX e. Numeral0 < e & e < d1 & e < d2";
+by (res_inst_tac [("x","(min d1 d2)/# 2")] exI 1);
by (asm_simp_tac (simpset() addsimps [min_def]) 1);
qed "hypreal_lbound_gt_zero";
@@ -560,7 +560,7 @@
by Auto_tac;
qed "hypreal_minus_equation";
-Goal "(x + - a = (#0::hypreal)) = (x=a)";
+Goal "(x + - a = (Numeral0::hypreal)) = (x=a)";
by (arith_tac 1);
qed "hypreal_add_minus_iff";
Addsimps [hypreal_add_minus_iff];
@@ -588,44 +588,44 @@
[hypreal_minus_less, hypreal_minus_le, hypreal_minus_equation]);
-(*** Simprules combining x+y and #0 ***)
+(*** Simprules combining x+y and Numeral0 ***)
-Goal "(x+y = (#0::hypreal)) = (y = -x)";
+Goal "(x+y = (Numeral0::hypreal)) = (y = -x)";
by Auto_tac;
qed "hypreal_add_eq_0_iff";
AddIffs [hypreal_add_eq_0_iff];
-Goal "(x+y < (#0::hypreal)) = (y < -x)";
+Goal "(x+y < (Numeral0::hypreal)) = (y < -x)";
by Auto_tac;
qed "hypreal_add_less_0_iff";
AddIffs [hypreal_add_less_0_iff];
-Goal "((#0::hypreal) < x+y) = (-x < y)";
+Goal "((Numeral0::hypreal) < x+y) = (-x < y)";
by Auto_tac;
qed "hypreal_0_less_add_iff";
AddIffs [hypreal_0_less_add_iff];
-Goal "(x+y <= (#0::hypreal)) = (y <= -x)";
+Goal "(x+y <= (Numeral0::hypreal)) = (y <= -x)";
by Auto_tac;
qed "hypreal_add_le_0_iff";
AddIffs [hypreal_add_le_0_iff];
-Goal "((#0::hypreal) <= x+y) = (-x <= y)";
+Goal "((Numeral0::hypreal) <= x+y) = (-x <= y)";
by Auto_tac;
qed "hypreal_0_le_add_iff";
AddIffs [hypreal_0_le_add_iff];
-(** Simprules combining x-y and #0; see also hypreal_less_iff_diff_less_0 etc
+(** Simprules combining x-y and Numeral0; see also hypreal_less_iff_diff_less_0 etc
in HyperBin
**)
-Goal "((#0::hypreal) < x-y) = (y < x)";
+Goal "((Numeral0::hypreal) < x-y) = (y < x)";
by Auto_tac;
qed "hypreal_0_less_diff_iff";
AddIffs [hypreal_0_less_diff_iff];
-Goal "((#0::hypreal) <= x-y) = (y <= x)";
+Goal "((Numeral0::hypreal) <= x-y) = (y <= x)";
by Auto_tac;
qed "hypreal_0_le_diff_iff";
AddIffs [hypreal_0_le_diff_iff];
@@ -644,11 +644,11 @@
(*** Density of the Hyperreals ***)
-Goal "x < y ==> x < (x+y) / (#2::hypreal)";
+Goal "x < y ==> x < (x+y) / (# 2::hypreal)";
by Auto_tac;
qed "hypreal_less_half_sum";
-Goal "x < y ==> (x+y)/(#2::hypreal) < y";
+Goal "x < y ==> (x+y)/(# 2::hypreal) < y";
by Auto_tac;
qed "hypreal_gt_half_sum";
@@ -657,7 +657,7 @@
qed "hypreal_dense";
-(*Replaces "inverse #nn" by #1/#nn *)
+(*Replaces "inverse #nn" by Numeral1/#nn *)
Addsimps [inst "x" "number_of ?w" hypreal_inverse_eq_divide];
--- a/src/HOL/Hyperreal/HyperBin.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Hyperreal/HyperBin.ML Fri Oct 05 21:52:39 2001 +0200
@@ -13,11 +13,11 @@
qed "hypreal_number_of";
Addsimps [hypreal_number_of];
-Goalw [hypreal_number_of_def] "(0::hypreal) = #0";
+Goalw [hypreal_number_of_def] "(0::hypreal) = Numeral0";
by (simp_tac (simpset() addsimps [hypreal_of_real_zero RS sym]) 1);
qed "zero_eq_numeral_0";
-Goalw [hypreal_number_of_def] "1hr = #1";
+Goalw [hypreal_number_of_def] "1hr = Numeral1";
by (simp_tac (simpset() addsimps [hypreal_of_real_one RS sym]) 1);
qed "one_eq_numeral_1";
@@ -57,18 +57,18 @@
qed "mult_hypreal_number_of";
Addsimps [mult_hypreal_number_of];
-Goal "(#2::hypreal) = #1 + #1";
+Goal "(# 2::hypreal) = Numeral1 + Numeral1";
by (Simp_tac 1);
val lemma = result();
(*For specialist use: NOT as default simprules*)
-Goal "#2 * z = (z+z::hypreal)";
+Goal "# 2 * z = (z+z::hypreal)";
by (simp_tac (simpset ()
addsimps [lemma, hypreal_add_mult_distrib,
one_eq_numeral_1 RS sym]) 1);
qed "hypreal_mult_2";
-Goal "z * #2 = (z+z::hypreal)";
+Goal "z * # 2 = (z+z::hypreal)";
by (stac hypreal_mult_commute 1 THEN rtac hypreal_mult_2 1);
qed "hypreal_mult_2_right";
@@ -107,11 +107,11 @@
(*** New versions of existing theorems involving 0, 1hr ***)
-Goal "- #1 = (#-1::hypreal)";
+Goal "- Numeral1 = (# -1::hypreal)";
by (Simp_tac 1);
qed "minus_numeral_one";
-(*Maps 0 to #0 and 1hr to #1 and -1hr to #-1*)
+(*Maps 0 to Numeral0 and 1hr to Numeral1 and -1hr to # -1*)
val hypreal_numeral_ss =
real_numeral_ss addsimps [zero_eq_numeral_0, one_eq_numeral_1,
minus_numeral_one];
@@ -176,15 +176,15 @@
(** Combining of literal coefficients in sums of products **)
-Goal "(x < y) = (x-y < (#0::hypreal))";
+Goal "(x < y) = (x-y < (Numeral0::hypreal))";
by (simp_tac (simpset() addsimps [hypreal_diff_less_eq]) 1);
qed "hypreal_less_iff_diff_less_0";
-Goal "(x = y) = (x-y = (#0::hypreal))";
+Goal "(x = y) = (x-y = (Numeral0::hypreal))";
by (simp_tac (simpset() addsimps [hypreal_diff_eq_eq]) 1);
qed "hypreal_eq_iff_diff_eq_0";
-Goal "(x <= y) = (x-y <= (#0::hypreal))";
+Goal "(x <= y) = (x-y <= (Numeral0::hypreal))";
by (simp_tac (simpset() addsimps [hypreal_diff_le_eq]) 1);
qed "hypreal_le_iff_diff_le_0";
@@ -242,14 +242,14 @@
hypreal_add_ac@rel_iff_rel_0_rls) 1);
qed "hypreal_le_add_iff2";
-Goal "(z::hypreal) * #-1 = -z";
+Goal "(z::hypreal) * # -1 = -z";
by (stac (minus_numeral_one RS sym) 1);
by (stac (hypreal_minus_mult_eq2 RS sym) 1);
by Auto_tac;
qed "hypreal_mult_minus_1_right";
Addsimps [hypreal_mult_minus_1_right];
-Goal "#-1 * (z::hypreal) = -z";
+Goal "# -1 * (z::hypreal) = -z";
by (simp_tac (simpset() addsimps [hypreal_mult_commute]) 1);
qed "hypreal_mult_minus_1";
Addsimps [hypreal_mult_minus_1];
@@ -275,7 +275,7 @@
val uminus_const = Const ("uminus", hyprealT --> hyprealT);
-(*Thus mk_sum[t] yields t+#0; longer sums don't have a trailing zero*)
+(*Thus mk_sum[t] yields t+Numeral0; longer sums don't have a trailing zero*)
fun mk_sum [] = zero
| mk_sum [t,u] = mk_plus (t, u)
| mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
@@ -335,7 +335,7 @@
handle TERM _ => find_first_coeff (t::past) u terms;
-(*Simplify #1*n and n*#1 to n*)
+(*Simplify Numeral1*n and n*Numeral1 to n*)
val add_0s = map rename_numerals
[hypreal_add_zero_left, hypreal_add_zero_right];
val mult_plus_1s = map rename_numerals
@@ -471,7 +471,7 @@
structure CombineNumeralsData =
struct
val add = op + : int*int -> int
- val mk_sum = long_mk_sum (*to work for e.g. #2*x + #3*x *)
+ val mk_sum = long_mk_sum (*to work for e.g. # 2*x + # 3*x *)
val dest_sum = dest_sum
val mk_coeff = mk_coeff
val dest_coeff = dest_coeff 1
@@ -530,34 +530,34 @@
set trace_simp;
fun test s = (Goal s, by (Simp_tac 1));
-test "l + #2 + #2 + #2 + (l + #2) + (oo + #2) = (uu::hypreal)";
-test "#2*u = (u::hypreal)";
-test "(i + j + #12 + (k::hypreal)) - #15 = y";
-test "(i + j + #12 + (k::hypreal)) - #5 = y";
+test "l + # 2 + # 2 + # 2 + (l + # 2) + (oo + # 2) = (uu::hypreal)";
+test "# 2*u = (u::hypreal)";
+test "(i + j + # 12 + (k::hypreal)) - # 15 = y";
+test "(i + j + # 12 + (k::hypreal)) - # 5 = y";
test "y - b < (b::hypreal)";
-test "y - (#3*b + c) < (b::hypreal) - #2*c";
+test "y - (# 3*b + c) < (b::hypreal) - # 2*c";
-test "(#2*x - (u*v) + y) - v*#3*u = (w::hypreal)";
-test "(#2*x*u*v + (u*v)*#4 + y) - v*u*#4 = (w::hypreal)";
-test "(#2*x*u*v + (u*v)*#4 + y) - v*u = (w::hypreal)";
-test "u*v - (x*u*v + (u*v)*#4 + y) = (w::hypreal)";
+test "(# 2*x - (u*v) + y) - v*# 3*u = (w::hypreal)";
+test "(# 2*x*u*v + (u*v)*# 4 + y) - v*u*# 4 = (w::hypreal)";
+test "(# 2*x*u*v + (u*v)*# 4 + y) - v*u = (w::hypreal)";
+test "u*v - (x*u*v + (u*v)*# 4 + y) = (w::hypreal)";
-test "(i + j + #12 + (k::hypreal)) = u + #15 + y";
-test "(i + j*#2 + #12 + (k::hypreal)) = j + #5 + y";
+test "(i + j + # 12 + (k::hypreal)) = u + # 15 + y";
+test "(i + j*# 2 + # 12 + (k::hypreal)) = j + # 5 + y";
-test "#2*y + #3*z + #6*w + #2*y + #3*z + #2*u = #2*y' + #3*z' + #6*w' + #2*y' + #3*z' + u + (vv::hypreal)";
+test "# 2*y + # 3*z + # 6*w + # 2*y + # 3*z + # 2*u = # 2*y' + # 3*z' + # 6*w' + # 2*y' + # 3*z' + u + (vv::hypreal)";
test "a + -(b+c) + b = (d::hypreal)";
test "a + -(b+c) - b = (d::hypreal)";
(*negative numerals*)
-test "(i + j + #-2 + (k::hypreal)) - (u + #5 + y) = zz";
-test "(i + j + #-3 + (k::hypreal)) < u + #5 + y";
-test "(i + j + #3 + (k::hypreal)) < u + #-6 + y";
-test "(i + j + #-12 + (k::hypreal)) - #15 = y";
-test "(i + j + #12 + (k::hypreal)) - #-15 = y";
-test "(i + j + #-12 + (k::hypreal)) - #-15 = y";
+test "(i + j + # -2 + (k::hypreal)) - (u + # 5 + y) = zz";
+test "(i + j + # -3 + (k::hypreal)) < u + # 5 + y";
+test "(i + j + # 3 + (k::hypreal)) < u + # -6 + y";
+test "(i + j + # -12 + (k::hypreal)) - # 15 = y";
+test "(i + j + # 12 + (k::hypreal)) - # -15 = y";
+test "(i + j + # -12 + (k::hypreal)) - # -15 = y";
*)
--- a/src/HOL/Hyperreal/HyperDef.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Hyperreal/HyperDef.ML Fri Oct 05 21:52:39 2001 +0200
@@ -304,7 +304,7 @@
by (asm_full_simp_tac (simpset() addsimps [hypreal_minus_minus]) 1);
qed "inj_hypreal_minus";
-Goalw [hypreal_zero_def] "-0 = (0::hypreal)";
+Goalw [hypreal_zero_def] "- 0 = (0::hypreal)";
by (simp_tac (simpset() addsimps [hypreal_minus]) 1);
qed "hypreal_minus_zero";
Addsimps [hypreal_minus_zero];
@@ -622,13 +622,13 @@
(**** multiplicative inverse on hypreal ****)
Goalw [congruent_def]
- "congruent hyprel (%X. hyprel``{%n. if X n = #0 then #0 else inverse(X n)})";
+ "congruent hyprel (%X. hyprel``{%n. if X n = Numeral0 then Numeral0 else inverse(X n)})";
by (Auto_tac THEN Ultra_tac 1);
qed "hypreal_inverse_congruent";
Goalw [hypreal_inverse_def]
"inverse (Abs_hypreal(hyprel``{%n. X n})) = \
-\ Abs_hypreal(hyprel `` {%n. if X n = #0 then #0 else inverse(X n)})";
+\ Abs_hypreal(hyprel `` {%n. if X n = Numeral0 then Numeral0 else inverse(X n)})";
by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1);
by (simp_tac (simpset() addsimps
[hyprel_in_hypreal RS Abs_hypreal_inverse,
@@ -840,8 +840,8 @@
Trichotomy of the hyperreals
--------------------------------------------------------------------------------*)
-Goalw [hyprel_def] "EX x. x: hyprel `` {%n. #0}";
-by (res_inst_tac [("x","%n. #0")] exI 1);
+Goalw [hyprel_def] "EX x. x: hyprel `` {%n. Numeral0}";
+by (res_inst_tac [("x","%n. Numeral0")] exI 1);
by (Step_tac 1);
by (auto_tac (claset() addSIs [FreeUltrafilterNat_Nat_set], simpset()));
qed "lemma_hyprel_0r_mem";
@@ -1101,22 +1101,22 @@
(*DON'T insert this or the next one as default simprules.
They are used in both orientations and anyway aren't the ones we finally
need, which would use binary literals.*)
-Goalw [hypreal_of_real_def,hypreal_one_def] "hypreal_of_real #1 = 1hr";
+Goalw [hypreal_of_real_def,hypreal_one_def] "hypreal_of_real Numeral1 = 1hr";
by (Step_tac 1);
qed "hypreal_of_real_one";
-Goalw [hypreal_of_real_def,hypreal_zero_def] "hypreal_of_real #0 = 0";
+Goalw [hypreal_of_real_def,hypreal_zero_def] "hypreal_of_real Numeral0 = 0";
by (Step_tac 1);
qed "hypreal_of_real_zero";
-Goal "(hypreal_of_real r = 0) = (r = #0)";
+Goal "(hypreal_of_real r = 0) = (r = Numeral0)";
by (auto_tac (claset() addIs [FreeUltrafilterNat_P],
simpset() addsimps [hypreal_of_real_def,
hypreal_zero_def,FreeUltrafilterNat_Nat_set]));
qed "hypreal_of_real_zero_iff";
Goal "hypreal_of_real (inverse r) = inverse (hypreal_of_real r)";
-by (case_tac "r=#0" 1);
+by (case_tac "r=Numeral0" 1);
by (asm_simp_tac (simpset() addsimps [REAL_DIVIDE_ZERO, INVERSE_ZERO,
HYPREAL_INVERSE_ZERO, hypreal_of_real_zero]) 1);
by (res_inst_tac [("c1","hypreal_of_real r")]
--- a/src/HOL/Hyperreal/HyperDef.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Hyperreal/HyperDef.thy Fri Oct 05 21:52:39 2001 +0200
@@ -35,10 +35,10 @@
defs
hypreal_zero_def
- "0 == Abs_hypreal(hyprel``{%n::nat. (#0::real)})"
+ "0 == Abs_hypreal(hyprel``{%n::nat. (Numeral0::real)})"
hypreal_one_def
- "1hr == Abs_hypreal(hyprel``{%n::nat. (#1::real)})"
+ "1hr == Abs_hypreal(hyprel``{%n::nat. (Numeral1::real)})"
hypreal_minus_def
"- P == Abs_hypreal(UN X: Rep_hypreal(P). hyprel``{%n::nat. - (X n)})"
@@ -48,7 +48,7 @@
hypreal_inverse_def
"inverse P == Abs_hypreal(UN X: Rep_hypreal(P).
- hyprel``{%n. if X n = #0 then #0 else inverse (X n)})"
+ hyprel``{%n. if X n = Numeral0 then Numeral0 else inverse (X n)})"
hypreal_divide_def
"P / Q::hypreal == P * inverse Q"
--- a/src/HOL/Hyperreal/HyperNat.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Hyperreal/HyperNat.ML Fri Oct 05 21:52:39 2001 +0200
@@ -682,7 +682,7 @@
Goalw [hypnat_one_def,hypnat_zero_def,hypnat_less_def]
"(0::hypnat) < 1hn";
by (res_inst_tac [("x","%n. 0")] exI 1);
-by (res_inst_tac [("x","%n. 1'")] exI 1);
+by (res_inst_tac [("x","%n. Suc 0")] exI 1);
by Auto_tac;
qed "hypnat_zero_less_one";
@@ -806,7 +806,7 @@
by Auto_tac;
qed "hypnat_of_nat_le_iff";
-Goalw [hypnat_of_nat_def,hypnat_one_def] "hypnat_of_nat 1' = 1hn";
+Goalw [hypnat_of_nat_def,hypnat_one_def] "hypnat_of_nat (Suc 0) = 1hn";
by (Simp_tac 1);
qed "hypnat_of_nat_one";
@@ -903,7 +903,7 @@
qed "SHNat_hypnat_of_nat";
Addsimps [SHNat_hypnat_of_nat];
-Goal "hypnat_of_nat 1' : Nats";
+Goal "hypnat_of_nat (Suc 0) : Nats";
by (Simp_tac 1);
qed "SHNat_hypnat_of_nat_one";
@@ -1246,14 +1246,14 @@
Addsimps [hypnat_of_nat_eq_cancel];
Goalw [hypnat_zero_def]
- "hypreal_of_hypnat 0 = #0";
+ "hypreal_of_hypnat 0 = Numeral0";
by (simp_tac (HOL_ss addsimps
[zero_eq_numeral_0 RS sym, hypreal_zero_def]) 1);
by (simp_tac (simpset() addsimps [hypreal_of_hypnat, real_of_nat_zero]) 1);
qed "hypreal_of_hypnat_zero";
Goalw [hypnat_one_def]
- "hypreal_of_hypnat 1hn = #1";
+ "hypreal_of_hypnat 1hn = Numeral1";
by (simp_tac (HOL_ss addsimps
[one_eq_numeral_1 RS sym, hypreal_one_def]) 1);
by (simp_tac (simpset() addsimps [hypreal_of_hypnat, real_of_nat_one]) 1);
@@ -1283,7 +1283,7 @@
qed "hypreal_of_hypnat_less_iff";
Addsimps [hypreal_of_hypnat_less_iff];
-Goal "(hypreal_of_hypnat N = #0) = (N = 0)";
+Goal "(hypreal_of_hypnat N = Numeral0) = (N = 0)";
by (simp_tac (simpset() addsimps [hypreal_of_hypnat_zero RS sym]) 1);
qed "hypreal_of_hypnat_eq_zero_iff";
Addsimps [hypreal_of_hypnat_eq_zero_iff];
--- a/src/HOL/Hyperreal/HyperOrd.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Hyperreal/HyperOrd.ML Fri Oct 05 21:52:39 2001 +0200
@@ -47,7 +47,7 @@
val eq_diff_eq = hypreal_eq_diff_eq
val eqI_rules = [hypreal_less_eqI, hypreal_eq_eqI, hypreal_le_eqI]
fun dest_eqI th =
- #1 (HOLogic.dest_bin "op =" HOLogic.boolT
+ #1 (HOLogic.dest_bin "op =" HOLogic.boolT
(HOLogic.dest_Trueprop (concl_of th)))
val diff_def = hypreal_diff_def
@@ -150,8 +150,8 @@
qed "hypreal_mult_less_zero";
Goalw [hypreal_one_def,hypreal_zero_def,hypreal_less_def] "0 < 1hr";
-by (res_inst_tac [("x","%n. #0")] exI 1);
-by (res_inst_tac [("x","%n. #1")] exI 1);
+by (res_inst_tac [("x","%n. Numeral0")] exI 1);
+by (res_inst_tac [("x","%n. Numeral1")] exI 1);
by (auto_tac (claset(),
simpset() addsimps [real_zero_less_one, FreeUltrafilterNat_Nat_set]));
qed "hypreal_zero_less_one";
--- a/src/HOL/Hyperreal/HyperPow.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Hyperreal/HyperPow.ML Fri Oct 05 21:52:39 2001 +0200
@@ -6,17 +6,17 @@
Exponentials on the hyperreals
*)
-Goal "(#0::hypreal) ^ (Suc n) = 0";
+Goal "(Numeral0::hypreal) ^ (Suc n) = 0";
by (Auto_tac);
qed "hrealpow_zero";
Addsimps [hrealpow_zero];
-Goal "r ~= (#0::hypreal) --> r ^ n ~= 0";
+Goal "r ~= (Numeral0::hypreal) --> r ^ n ~= 0";
by (induct_tac "n" 1);
by Auto_tac;
qed_spec_mp "hrealpow_not_zero";
-Goal "r ~= (#0::hypreal) --> inverse(r ^ n) = (inverse r) ^ n";
+Goal "r ~= (Numeral0::hypreal) --> inverse(r ^ n) = (inverse r) ^ n";
by (induct_tac "n" 1);
by (Auto_tac);
by (forw_inst_tac [("n","n")] hrealpow_not_zero 1);
@@ -33,49 +33,49 @@
by (auto_tac (claset(), simpset() addsimps hypreal_mult_ac));
qed "hrealpow_add";
-Goal "(r::hypreal) ^ 1' = r";
+Goal "(r::hypreal) ^ Suc 0 = r";
by (Simp_tac 1);
qed "hrealpow_one";
Addsimps [hrealpow_one];
-Goal "(r::hypreal) ^ 2 = r * r";
+Goal "(r::hypreal) ^ Suc (Suc 0) = r * r";
by (Simp_tac 1);
qed "hrealpow_two";
-Goal "(#0::hypreal) <= r --> #0 <= r ^ n";
+Goal "(Numeral0::hypreal) <= r --> Numeral0 <= r ^ n";
by (induct_tac "n" 1);
by (auto_tac (claset(), simpset() addsimps [hypreal_0_le_mult_iff]));
qed_spec_mp "hrealpow_ge_zero";
-Goal "(#0::hypreal) < r --> #0 < r ^ n";
+Goal "(Numeral0::hypreal) < r --> Numeral0 < r ^ n";
by (induct_tac "n" 1);
by (auto_tac (claset(), simpset() addsimps [hypreal_0_less_mult_iff]));
qed_spec_mp "hrealpow_gt_zero";
-Goal "x <= y & (#0::hypreal) < x --> x ^ n <= y ^ n";
+Goal "x <= y & (Numeral0::hypreal) < x --> x ^ n <= y ^ n";
by (induct_tac "n" 1);
by (auto_tac (claset() addSIs [hypreal_mult_le_mono], simpset()));
by (asm_simp_tac (simpset() addsimps [hrealpow_ge_zero]) 1);
qed_spec_mp "hrealpow_le";
-Goal "x < y & (#0::hypreal) < x & 0 < n --> x ^ n < y ^ n";
+Goal "x < y & (Numeral0::hypreal) < x & 0 < n --> x ^ n < y ^ n";
by (induct_tac "n" 1);
by (auto_tac (claset() addIs [hypreal_mult_less_mono,gr0I],
simpset() addsimps [hrealpow_gt_zero]));
qed "hrealpow_less";
-Goal "#1 ^ n = (#1::hypreal)";
+Goal "Numeral1 ^ n = (Numeral1::hypreal)";
by (induct_tac "n" 1);
by (Auto_tac);
qed "hrealpow_eq_one";
Addsimps [hrealpow_eq_one];
-Goal "abs(-(#1 ^ n)) = (#1::hypreal)";
+Goal "abs(-(Numeral1 ^ n)) = (Numeral1::hypreal)";
by Auto_tac;
qed "hrabs_minus_hrealpow_one";
Addsimps [hrabs_minus_hrealpow_one];
-Goal "abs(#-1 ^ n) = (#1::hypreal)";
+Goal "abs(# -1 ^ n) = (Numeral1::hypreal)";
by (induct_tac "n" 1);
by Auto_tac;
qed "hrabs_hrealpow_minus_one";
@@ -86,61 +86,61 @@
by (auto_tac (claset(), simpset() addsimps hypreal_mult_ac));
qed "hrealpow_mult";
-Goal "(#0::hypreal) <= r ^ 2";
+Goal "(Numeral0::hypreal) <= r ^Suc (Suc 0)";
by (auto_tac (claset(), simpset() addsimps [hypreal_0_le_mult_iff]));
qed "hrealpow_two_le";
Addsimps [hrealpow_two_le];
-Goal "(#0::hypreal) <= u ^ 2 + v ^ 2";
+Goal "(Numeral0::hypreal) <= u ^ Suc (Suc 0) + v ^ Suc (Suc 0)";
by (simp_tac (HOL_ss addsimps [hrealpow_two_le,
rename_numerals hypreal_le_add_order]) 1);
qed "hrealpow_two_le_add_order";
Addsimps [hrealpow_two_le_add_order];
-Goal "(#0::hypreal) <= u ^ 2 + v ^ 2 + w ^ 2";
+Goal "(Numeral0::hypreal) <= u ^ Suc (Suc 0) + v ^ Suc (Suc 0) + w ^ Suc (Suc 0)";
by (simp_tac (HOL_ss addsimps [hrealpow_two_le,
rename_numerals hypreal_le_add_order]) 1);
qed "hrealpow_two_le_add_order2";
Addsimps [hrealpow_two_le_add_order2];
-Goal "(x ^ 2 + y ^ 2 + z ^ 2 = (#0::hypreal)) = (x = #0 & y = #0 & z = #0)";
+Goal "(x ^ Suc (Suc 0) + y ^ Suc (Suc 0) + z ^ Suc (Suc 0) = (Numeral0::hypreal)) = (x = Numeral0 & y = Numeral0 & z = Numeral0)";
by (simp_tac (HOL_ss addsimps
[rename_numerals hypreal_three_squares_add_zero_iff, hrealpow_two]) 1);
qed "hrealpow_three_squares_add_zero_iff";
Addsimps [hrealpow_three_squares_add_zero_iff];
-Goal "abs(x ^ 2) = (x::hypreal) ^ 2";
+Goal "abs(x ^ Suc (Suc 0)) = (x::hypreal) ^ Suc (Suc 0)";
by (auto_tac (claset(),
simpset() addsimps [hrabs_def, hypreal_0_le_mult_iff]));
qed "hrabs_hrealpow_two";
Addsimps [hrabs_hrealpow_two];
-Goal "abs(x) ^ 2 = (x::hypreal) ^ 2";
+Goal "abs(x) ^ Suc (Suc 0) = (x::hypreal) ^ Suc (Suc 0)";
by (simp_tac (simpset() addsimps [hrealpow_hrabs, hrabs_eqI1]
delsimps [hpowr_Suc]) 1);
qed "hrealpow_two_hrabs";
Addsimps [hrealpow_two_hrabs];
-Goal "(#1::hypreal) < r ==> #1 < r ^ 2";
+Goal "(Numeral1::hypreal) < r ==> Numeral1 < r ^ Suc (Suc 0)";
by (auto_tac (claset(), simpset() addsimps [hrealpow_two]));
-by (res_inst_tac [("y","#1*#1")] order_le_less_trans 1);
+by (res_inst_tac [("y","Numeral1*Numeral1")] order_le_less_trans 1);
by (rtac hypreal_mult_less_mono 2);
by Auto_tac;
qed "hrealpow_two_gt_one";
-Goal "(#1::hypreal) <= r ==> #1 <= r ^ 2";
+Goal "(Numeral1::hypreal) <= r ==> Numeral1 <= r ^ Suc (Suc 0)";
by (etac (order_le_imp_less_or_eq RS disjE) 1);
by (etac (hrealpow_two_gt_one RS order_less_imp_le) 1);
by Auto_tac;
qed "hrealpow_two_ge_one";
-Goal "(#1::hypreal) <= #2 ^ n";
-by (res_inst_tac [("y","#1 ^ n")] order_trans 1);
+Goal "(Numeral1::hypreal) <= # 2 ^ n";
+by (res_inst_tac [("y","Numeral1 ^ n")] order_trans 1);
by (rtac hrealpow_le 2);
by Auto_tac;
qed "two_hrealpow_ge_one";
-Goal "hypreal_of_nat n < #2 ^ n";
+Goal "hypreal_of_nat n < # 2 ^ n";
by (induct_tac "n" 1);
by (auto_tac (claset(),
simpset() addsimps [hypreal_of_nat_Suc, hypreal_add_mult_distrib]));
@@ -149,34 +149,34 @@
qed "two_hrealpow_gt";
Addsimps [two_hrealpow_gt,two_hrealpow_ge_one];
-Goal "#-1 ^ (#2*n) = (#1::hypreal)";
+Goal "# -1 ^ (# 2*n) = (Numeral1::hypreal)";
by (induct_tac "n" 1);
by (Auto_tac);
qed "hrealpow_minus_one";
-Goal "n+n = (#2*n::nat)";
+Goal "n+n = (# 2*n::nat)";
by Auto_tac;
qed "double_lemma";
(*ugh: need to get rid fo the n+n*)
-Goal "#-1 ^ (n + n) = (#1::hypreal)";
+Goal "# -1 ^ (n + n) = (Numeral1::hypreal)";
by (auto_tac (claset(),
simpset() addsimps [double_lemma, hrealpow_minus_one]));
qed "hrealpow_minus_one2";
Addsimps [hrealpow_minus_one2];
-Goal "(-(x::hypreal)) ^ 2 = x ^ 2";
+Goal "(-(x::hypreal)) ^ Suc (Suc 0) = x ^ Suc (Suc 0)";
by (Auto_tac);
qed "hrealpow_minus_two";
Addsimps [hrealpow_minus_two];
-Goal "(#0::hypreal) < r & r < #1 --> r ^ Suc n < r ^ n";
+Goal "(Numeral0::hypreal) < r & r < Numeral1 --> r ^ Suc n < r ^ n";
by (induct_tac "n" 1);
by (auto_tac (claset(),
simpset() addsimps [hypreal_mult_less_mono2]));
qed_spec_mp "hrealpow_Suc_less";
-Goal "(#0::hypreal) <= r & r < #1 --> r ^ Suc n <= r ^ n";
+Goal "(Numeral0::hypreal) <= r & r < Numeral1 --> r ^ Suc n <= r ^ n";
by (induct_tac "n" 1);
by (auto_tac (claset() addIs [order_less_imp_le]
addSDs [order_le_imp_less_or_eq,hrealpow_Suc_less],
@@ -191,8 +191,8 @@
one_eq_numeral_1 RS sym]));
qed "hrealpow";
-Goal "(x + (y::hypreal)) ^ 2 = \
-\ x ^ 2 + y ^ 2 + (hypreal_of_nat 2)*x*y";
+Goal "(x + (y::hypreal)) ^ Suc (Suc 0) = \
+\ x ^ Suc (Suc 0) + y ^ Suc (Suc 0) + (hypreal_of_nat (Suc (Suc 0)))*x*y";
by (simp_tac (simpset() addsimps
[hypreal_add_mult_distrib2, hypreal_add_mult_distrib,
hypreal_of_nat_zero, hypreal_of_nat_Suc]) 1);
@@ -204,7 +204,7 @@
property for the real rather than prove it directly
using induction: proof is much simpler this way!
---------------------------------------------------------------*)
-Goal "[|(#0::hypreal) <= x; #0 <= y;x ^ Suc n <= y ^ Suc n |] ==> x <= y";
+Goal "[|(Numeral0::hypreal) <= x; Numeral0 <= y;x ^ Suc n <= y ^ Suc n |] ==> x <= y";
by (full_simp_tac (simpset() addsimps [rename_numerals hypreal_zero_def]) 1);
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
@@ -241,14 +241,14 @@
by (Fuf_tac 1);
qed "hyperpow";
-Goalw [hypnat_one_def] "(#0::hypreal) pow (n + 1hn) = #0";
+Goalw [hypnat_one_def] "(Numeral0::hypreal) pow (n + 1hn) = Numeral0";
by (simp_tac (simpset() addsimps [rename_numerals hypreal_zero_def]) 1);
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
by (auto_tac (claset(), simpset() addsimps [hyperpow,hypnat_add]));
qed "hyperpow_zero";
Addsimps [hyperpow_zero];
-Goal "r ~= (#0::hypreal) --> r pow n ~= #0";
+Goal "r ~= (Numeral0::hypreal) --> r pow n ~= Numeral0";
by (simp_tac (simpset() addsimps [rename_numerals hypreal_zero_def]) 1);
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
@@ -258,7 +258,7 @@
simpset()) 1);
qed_spec_mp "hyperpow_not_zero";
-Goal "r ~= (#0::hypreal) --> inverse(r pow n) = (inverse r) pow n";
+Goal "r ~= (Numeral0::hypreal) --> inverse(r pow n) = (inverse r) pow n";
by (simp_tac (simpset() addsimps [rename_numerals hypreal_zero_def]) 1);
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
@@ -298,7 +298,7 @@
simpset() addsimps [hyperpow,hypnat_add, hypreal_mult]));
qed "hyperpow_two";
-Goal "(#0::hypreal) < r --> #0 < r pow n";
+Goal "(Numeral0::hypreal) < r --> Numeral0 < r pow n";
by (simp_tac (simpset() addsimps [rename_numerals hypreal_zero_def]) 1);
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
@@ -306,7 +306,7 @@
simpset() addsimps [hyperpow,hypreal_less, hypreal_le]));
qed_spec_mp "hyperpow_gt_zero";
-Goal "(#0::hypreal) <= r --> #0 <= r pow n";
+Goal "(Numeral0::hypreal) <= r --> Numeral0 <= r pow n";
by (simp_tac (simpset() addsimps [rename_numerals hypreal_zero_def]) 1);
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
@@ -314,7 +314,7 @@
simpset() addsimps [hyperpow,hypreal_le]));
qed "hyperpow_ge_zero";
-Goal "(#0::hypreal) < x & x <= y --> x pow n <= y pow n";
+Goal "(Numeral0::hypreal) < x & x <= y --> x pow n <= y pow n";
by (full_simp_tac (simpset() addsimps [rename_numerals hypreal_zero_def]) 1);
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
@@ -326,22 +326,22 @@
by (auto_tac (claset() addIs [realpow_le], simpset()));
qed_spec_mp "hyperpow_le";
-Goal "#1 pow n = (#1::hypreal)";
+Goal "Numeral1 pow n = (Numeral1::hypreal)";
by (simp_tac (HOL_ss addsimps [one_eq_numeral_1 RS sym, hypreal_one_def]) 1);
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
by (auto_tac (claset(), simpset() addsimps [hyperpow]));
qed "hyperpow_eq_one";
Addsimps [hyperpow_eq_one];
-Goal "abs(-(#1 pow n)) = (#1::hypreal)";
+Goal "abs(-(Numeral1 pow n)) = (Numeral1::hypreal)";
by (simp_tac (HOL_ss addsimps [one_eq_numeral_1 RS sym, hypreal_one_def]) 1);
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
by (auto_tac (claset(), simpset() addsimps [hyperpow,hypreal_hrabs]));
qed "hrabs_minus_hyperpow_one";
Addsimps [hrabs_minus_hyperpow_one];
-Goal "abs(#-1 pow n) = (#1::hypreal)";
-by (subgoal_tac "abs((-1hr) pow n) = 1hr" 1);
+Goal "abs(# -1 pow n) = (Numeral1::hypreal)";
+by (subgoal_tac "abs((- 1hr) pow n) = 1hr" 1);
by (Asm_full_simp_tac 1);
by (simp_tac (HOL_ss addsimps [hypreal_one_def]) 1);
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
@@ -358,7 +358,7 @@
simpset() addsimps [hyperpow, hypreal_mult,realpow_mult]));
qed "hyperpow_mult";
-Goal "(#0::hypreal) <= r pow (1hn + 1hn)";
+Goal "(Numeral0::hypreal) <= r pow (1hn + 1hn)";
by (auto_tac (claset(),
simpset() addsimps [hyperpow_two, hypreal_0_le_mult_iff]));
qed "hyperpow_two_le";
@@ -375,21 +375,21 @@
Addsimps [hyperpow_two_hrabs];
(*? very similar to hrealpow_two_gt_one *)
-Goal "(#1::hypreal) < r ==> #1 < r pow (1hn + 1hn)";
+Goal "(Numeral1::hypreal) < r ==> Numeral1 < r pow (1hn + 1hn)";
by (auto_tac (claset(), simpset() addsimps [hyperpow_two]));
-by (res_inst_tac [("y","#1*#1")] order_le_less_trans 1);
+by (res_inst_tac [("y","Numeral1*Numeral1")] order_le_less_trans 1);
by (rtac hypreal_mult_less_mono 2);
by Auto_tac;
qed "hyperpow_two_gt_one";
-Goal "(#1::hypreal) <= r ==> #1 <= r pow (1hn + 1hn)";
+Goal "(Numeral1::hypreal) <= r ==> Numeral1 <= r pow (1hn + 1hn)";
by (auto_tac (claset() addSDs [order_le_imp_less_or_eq]
addIs [hyperpow_two_gt_one,order_less_imp_le],
simpset()));
qed "hyperpow_two_ge_one";
-Goal "(#1::hypreal) <= #2 pow n";
-by (res_inst_tac [("y","#1 pow n")] order_trans 1);
+Goal "(Numeral1::hypreal) <= # 2 pow n";
+by (res_inst_tac [("y","Numeral1 pow n")] order_trans 1);
by (rtac hyperpow_le 2);
by Auto_tac;
qed "two_hyperpow_ge_one";
@@ -397,7 +397,7 @@
Addsimps [simplify (simpset()) realpow_minus_one];
-Goal "#-1 pow ((1hn + 1hn)*n) = (#1::hypreal)";
+Goal "# -1 pow ((1hn + 1hn)*n) = (Numeral1::hypreal)";
by (subgoal_tac "(-(1hr)) pow ((1hn + 1hn)*n) = 1hr" 1);
by (Asm_full_simp_tac 1);
by (simp_tac (HOL_ss addsimps [hypreal_one_def]) 1);
@@ -409,7 +409,7 @@
Addsimps [hyperpow_minus_one2];
Goalw [hypnat_one_def]
- "(#0::hypreal) < r & r < #1 --> r pow (n + 1hn) < r pow n";
+ "(Numeral0::hypreal) < r & r < Numeral1 --> r pow (n + 1hn) < r pow n";
by (full_simp_tac
(HOL_ss addsimps [zero_eq_numeral_0 RS sym, hypreal_zero_def,
one_eq_numeral_1 RS sym, hypreal_one_def]) 1);
@@ -421,7 +421,7 @@
qed_spec_mp "hyperpow_Suc_less";
Goalw [hypnat_one_def]
- "#0 <= r & r < (#1::hypreal) --> r pow (n + 1hn) <= r pow n";
+ "Numeral0 <= r & r < (Numeral1::hypreal) --> r pow (n + 1hn) <= r pow n";
by (full_simp_tac
(HOL_ss addsimps [zero_eq_numeral_0 RS sym, hypreal_zero_def,
one_eq_numeral_1 RS sym, hypreal_one_def]) 1);
@@ -434,7 +434,7 @@
qed_spec_mp "hyperpow_Suc_le";
Goalw [hypnat_one_def]
- "(#0::hypreal) <= r & r < #1 & n < N --> r pow N <= r pow n";
+ "(Numeral0::hypreal) <= r & r < Numeral1 & n < N --> r pow N <= r pow n";
by (full_simp_tac
(HOL_ss addsimps [zero_eq_numeral_0 RS sym, hypreal_zero_def,
one_eq_numeral_1 RS sym, hypreal_one_def]) 1);
@@ -449,12 +449,12 @@
simpset()));
qed_spec_mp "hyperpow_less_le";
-Goal "[| (#0::hypreal) <= r; r < #1 |] \
+Goal "[| (Numeral0::hypreal) <= r; r < Numeral1 |] \
\ ==> ALL N n. n < N --> r pow N <= r pow n";
by (blast_tac (claset() addSIs [hyperpow_less_le]) 1);
qed "hyperpow_less_le2";
-Goal "[| #0 <= r; r < (#1::hypreal); N : HNatInfinite |] \
+Goal "[| Numeral0 <= r; r < (Numeral1::hypreal); N : HNatInfinite |] \
\ ==> ALL n: Nats. r pow N <= r pow n";
by (auto_tac (claset() addSIs [hyperpow_less_le],
simpset() addsimps [HNatInfinite_iff]));
@@ -471,23 +471,23 @@
qed "hyperpow_SReal";
Addsimps [hyperpow_SReal];
-Goal "N : HNatInfinite ==> (#0::hypreal) pow N = 0";
+Goal "N : HNatInfinite ==> (Numeral0::hypreal) pow N = 0";
by (dtac HNatInfinite_is_Suc 1);
by (Auto_tac);
qed "hyperpow_zero_HNatInfinite";
Addsimps [hyperpow_zero_HNatInfinite];
-Goal "[| (#0::hypreal) <= r; r < #1; n <= N |] ==> r pow N <= r pow n";
+Goal "[| (Numeral0::hypreal) <= r; r < Numeral1; n <= N |] ==> r pow N <= r pow n";
by (dres_inst_tac [("y","N")] hypnat_le_imp_less_or_eq 1);
by (auto_tac (claset() addIs [hyperpow_less_le], simpset()));
qed "hyperpow_le_le";
-Goal "[| (#0::hypreal) < r; r < #1 |] ==> r pow (n + 1hn) <= r";
+Goal "[| (Numeral0::hypreal) < r; r < Numeral1 |] ==> r pow (n + 1hn) <= r";
by (dres_inst_tac [("n","1hn")] (order_less_imp_le RS hyperpow_le_le) 1);
by (Auto_tac);
qed "hyperpow_Suc_le_self";
-Goal "[| (#0::hypreal) <= r; r < #1 |] ==> r pow (n + 1hn) <= r";
+Goal "[| (Numeral0::hypreal) <= r; r < Numeral1 |] ==> r pow (n + 1hn) <= r";
by (dres_inst_tac [("n","1hn")] hyperpow_le_le 1);
by (Auto_tac);
qed "hyperpow_Suc_le_self2";
--- a/src/HOL/Hyperreal/Lim.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Hyperreal/Lim.ML Fri Oct 05 21:52:39 2001 +0200
@@ -29,7 +29,7 @@
Goalw [LIM_def]
"[| f -- x --> l; g -- x --> m |] ==> (%x. f(x) + g(x)) -- x --> (l + m)";
by (Clarify_tac 1);
-by (REPEAT(dres_inst_tac [("x","r/#2")] spec 1));
+by (REPEAT(dres_inst_tac [("x","r/# 2")] spec 1));
by (Asm_full_simp_tac 1);
by (Clarify_tac 1);
by (res_inst_tac [("R1.0","s"),("R2.0","sa")]
@@ -65,7 +65,7 @@
(*----------------------------------------------
LIM_zero
----------------------------------------------*)
-Goal "f -- a --> l ==> (%x. f(x) + -l) -- a --> #0";
+Goal "f -- a --> l ==> (%x. f(x) + -l) -- a --> Numeral0";
by (res_inst_tac [("z1","l")] (rename_numerals (real_add_minus RS subst)) 1);
by (rtac LIM_add_minus 1 THEN Auto_tac);
qed "LIM_zero";
@@ -73,8 +73,8 @@
(*--------------------------
Limit not zero
--------------------------*)
-Goalw [LIM_def] "k \\<noteq> #0 ==> ~ ((%x. k) -- x --> #0)";
-by (res_inst_tac [("R1.0","k"),("R2.0","#0")] real_linear_less2 1);
+Goalw [LIM_def] "k \\<noteq> Numeral0 ==> ~ ((%x. k) -- x --> Numeral0)";
+by (res_inst_tac [("R1.0","k"),("R2.0","Numeral0")] real_linear_less2 1);
by (auto_tac (claset(), simpset() addsimps [real_abs_def]));
by (res_inst_tac [("x","-k")] exI 1);
by (res_inst_tac [("x","k")] exI 2);
@@ -85,7 +85,7 @@
by Auto_tac;
qed "LIM_not_zero";
-(* [| k \\<noteq> #0; (%x. k) -- x --> #0 |] ==> R *)
+(* [| k \\<noteq> Numeral0; (%x. k) -- x --> Numeral0 |] ==> R *)
bind_thm("LIM_not_zeroE", LIM_not_zero RS notE);
Goal "(%x. k) -- x --> L ==> k = L";
@@ -108,9 +108,9 @@
LIM_mult_zero
-------------*)
Goalw [LIM_def]
- "[| f -- x --> #0; g -- x --> #0 |] ==> (%x. f(x)*g(x)) -- x --> #0";
+ "[| f -- x --> Numeral0; g -- x --> Numeral0 |] ==> (%x. f(x)*g(x)) -- x --> Numeral0";
by Safe_tac;
-by (dres_inst_tac [("x","#1")] spec 1);
+by (dres_inst_tac [("x","Numeral1")] spec 1);
by (dres_inst_tac [("x","r")] spec 1);
by (cut_facts_tac [real_zero_less_one] 1);
by (asm_full_simp_tac (simpset() addsimps
@@ -146,7 +146,7 @@
by (auto_tac (claset(), simpset() addsimps [real_add_minus_iff]));
qed "LIM_equal";
-Goal "[| (%x. f(x) + -g(x)) -- a --> #0; g -- a --> l |] \
+Goal "[| (%x. f(x) + -g(x)) -- a --> Numeral0; g -- a --> l |] \
\ ==> f -- a --> l";
by (dtac LIM_add 1 THEN assume_tac 1);
by (auto_tac (claset(), simpset() addsimps [real_add_assoc]));
@@ -181,7 +181,7 @@
Limit: NS definition ==> standard definition
---------------------------------------------------------------------*)
-Goal "\\<forall>s. #0 < s --> (\\<exists>xa. xa \\<noteq> x & \
+Goal "\\<forall>s. Numeral0 < s --> (\\<exists>xa. xa \\<noteq> x & \
\ abs (xa + - x) < s & r \\<le> abs (f xa + -L)) \
\ ==> \\<forall>n::nat. \\<exists>xa. xa \\<noteq> x & \
\ abs(xa + -x) < inverse(real(Suc n)) & r \\<le> abs(f xa + -L)";
@@ -191,7 +191,7 @@
by Auto_tac;
val lemma_LIM = result();
-Goal "\\<forall>s. #0 < s --> (\\<exists>xa. xa \\<noteq> x & \
+Goal "\\<forall>s. Numeral0 < s --> (\\<exists>xa. xa \\<noteq> x & \
\ abs (xa + - x) < s & r \\<le> abs (f xa + -L)) \
\ ==> \\<exists>X. \\<forall>n::nat. X n \\<noteq> x & \
\ abs(X n + -x) < inverse(real(Suc n)) & r \\<le> abs(f (X n) + -L)";
@@ -320,7 +320,7 @@
NSLIM_inverse
-----------------------------*)
Goalw [NSLIM_def]
- "[| f -- a --NS> L; L \\<noteq> #0 |] \
+ "[| f -- a --NS> L; L \\<noteq> Numeral0 |] \
\ ==> (%x. inverse(f(x))) -- a --NS> (inverse L)";
by (Clarify_tac 1);
by (dtac spec 1);
@@ -329,28 +329,28 @@
qed "NSLIM_inverse";
Goal "[| f -- a --> L; \
-\ L \\<noteq> #0 |] ==> (%x. inverse(f(x))) -- a --> (inverse L)";
+\ L \\<noteq> Numeral0 |] ==> (%x. inverse(f(x))) -- a --> (inverse L)";
by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_inverse]) 1);
qed "LIM_inverse";
(*------------------------------
NSLIM_zero
------------------------------*)
-Goal "f -- a --NS> l ==> (%x. f(x) + -l) -- a --NS> #0";
+Goal "f -- a --NS> l ==> (%x. f(x) + -l) -- a --NS> Numeral0";
by (res_inst_tac [("z1","l")] (rename_numerals (real_add_minus RS subst)) 1);
by (rtac NSLIM_add_minus 1 THEN Auto_tac);
qed "NSLIM_zero";
-Goal "f -- a --> l ==> (%x. f(x) + -l) -- a --> #0";
+Goal "f -- a --> l ==> (%x. f(x) + -l) -- a --> Numeral0";
by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_zero]) 1);
qed "LIM_zero2";
-Goal "(%x. f(x) - l) -- x --NS> #0 ==> f -- x --NS> l";
+Goal "(%x. f(x) - l) -- x --NS> Numeral0 ==> f -- x --NS> l";
by (dres_inst_tac [("g","%x. l"),("m","l")] NSLIM_add 1);
by (auto_tac (claset(),simpset() addsimps [real_diff_def, real_add_assoc]));
qed "NSLIM_zero_cancel";
-Goal "(%x. f(x) - l) -- x --> #0 ==> f -- x --> l";
+Goal "(%x. f(x) - l) -- x --> Numeral0 ==> f -- x --> l";
by (dres_inst_tac [("g","%x. l"),("m","l")] LIM_add 1);
by (auto_tac (claset(),simpset() addsimps [real_diff_def, real_add_assoc]));
qed "LIM_zero_cancel";
@@ -359,17 +359,17 @@
(*--------------------------
NSLIM_not_zero
--------------------------*)
-Goalw [NSLIM_def] "k \\<noteq> #0 ==> ~ ((%x. k) -- x --NS> #0)";
+Goalw [NSLIM_def] "k \\<noteq> Numeral0 ==> ~ ((%x. k) -- x --NS> Numeral0)";
by Auto_tac;
by (res_inst_tac [("x","hypreal_of_real x + epsilon")] exI 1);
by (auto_tac (claset() addIs [Infinitesimal_add_approx_self RS approx_sym],
simpset() addsimps [rename_numerals hypreal_epsilon_not_zero]));
qed "NSLIM_not_zero";
-(* [| k \\<noteq> #0; (%x. k) -- x --NS> #0 |] ==> R *)
+(* [| k \\<noteq> Numeral0; (%x. k) -- x --NS> Numeral0 |] ==> R *)
bind_thm("NSLIM_not_zeroE", NSLIM_not_zero RS notE);
-Goal "k \\<noteq> #0 ==> ~ ((%x. k) -- x --> #0)";
+Goal "k \\<noteq> Numeral0 ==> ~ ((%x. k) -- x --> Numeral0)";
by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_not_zero]) 1);
qed "LIM_not_zero2";
@@ -405,16 +405,16 @@
(*--------------------
NSLIM_mult_zero
--------------------*)
-Goal "[| f -- x --NS> #0; g -- x --NS> #0 |] \
-\ ==> (%x. f(x)*g(x)) -- x --NS> #0";
+Goal "[| f -- x --NS> Numeral0; g -- x --NS> Numeral0 |] \
+\ ==> (%x. f(x)*g(x)) -- x --NS> Numeral0";
by (dtac NSLIM_mult 1 THEN Auto_tac);
qed "NSLIM_mult_zero";
(* we can use the corresponding thm LIM_mult2 *)
(* for standard definition of limit *)
-Goal "[| f -- x --> #0; g -- x --> #0 |] \
-\ ==> (%x. f(x)*g(x)) -- x --> #0";
+Goal "[| f -- x --> Numeral0; g -- x --> Numeral0 |] \
+\ ==> (%x. f(x)*g(x)) -- x --> Numeral0";
by (dtac LIM_mult2 1 THEN Auto_tac);
qed "LIM_mult_zero2";
@@ -499,7 +499,7 @@
--------------------------------------------------------------------------*)
(* Prove equivalence between NS limits - *)
(* seems easier than using standard def *)
-Goalw [NSLIM_def] "(f -- a --NS> L) = ((%h. f(a + h)) -- #0 --NS> L)";
+Goalw [NSLIM_def] "(f -- a --NS> L) = ((%h. f(a + h)) -- Numeral0 --NS> L)";
by (auto_tac (claset(),simpset() addsimps [hypreal_of_real_zero]));
by (dres_inst_tac [("x","hypreal_of_real a + x")] spec 1);
by (dres_inst_tac [("x","-hypreal_of_real a + x")] spec 2);
@@ -516,15 +516,15 @@
hypreal_add, real_add_assoc, approx_refl, hypreal_zero_def]));
qed "NSLIM_h_iff";
-Goal "(f -- a --NS> f a) = ((%h. f(a + h)) -- #0 --NS> f a)";
+Goal "(f -- a --NS> f a) = ((%h. f(a + h)) -- Numeral0 --NS> f a)";
by (rtac NSLIM_h_iff 1);
qed "NSLIM_isCont_iff";
-Goal "(f -- a --> f a) = ((%h. f(a + h)) -- #0 --> f(a))";
+Goal "(f -- a --> f a) = ((%h. f(a + h)) -- Numeral0 --> f(a))";
by (simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_isCont_iff]) 1);
qed "LIM_isCont_iff";
-Goalw [isCont_def] "(isCont f x) = ((%h. f(x + h)) -- #0 --> f(x))";
+Goalw [isCont_def] "(isCont f x) = ((%h. f(x + h)) -- Numeral0 --> f(x))";
by (simp_tac (simpset() addsimps [LIM_isCont_iff]) 1);
qed "isCont_iff";
@@ -574,11 +574,11 @@
qed "isCont_minus";
Goalw [isCont_def]
- "[| isCont f x; f x \\<noteq> #0 |] ==> isCont (%x. inverse (f x)) x";
+ "[| isCont f x; f x \\<noteq> Numeral0 |] ==> isCont (%x. inverse (f x)) x";
by (blast_tac (claset() addIs [LIM_inverse]) 1);
qed "isCont_inverse";
-Goal "[| isNSCont f x; f x \\<noteq> #0 |] ==> isNSCont (%x. inverse (f x)) x";
+Goal "[| isNSCont f x; f x \\<noteq> Numeral0 |] ==> isNSCont (%x. inverse (f x)) x";
by (auto_tac (claset() addIs [isCont_inverse],simpset() addsimps
[isNSCont_isCont_iff]));
qed "isNSCont_inverse";
@@ -690,7 +690,7 @@
by (Ultra_tac 1);
qed "isUCont_isNSUCont";
-Goal "\\<forall>s. #0 < s --> (\\<exists>z y. abs (z + - y) < s & r \\<le> abs (f z + -f y)) \
+Goal "\\<forall>s. Numeral0 < s --> (\\<exists>z y. abs (z + - y) < s & r \\<le> abs (f z + -f y)) \
\ ==> \\<forall>n::nat. \\<exists>z y. \
\ abs(z + -y) < inverse(real(Suc n)) & \
\ r \\<le> abs(f z + -f y)";
@@ -700,7 +700,7 @@
by Auto_tac;
val lemma_LIMu = result();
-Goal "\\<forall>s. #0 < s --> (\\<exists>z y. abs (z + - y) < s & r \\<le> abs (f z + -f y)) \
+Goal "\\<forall>s. Numeral0 < s --> (\\<exists>z y. abs (z + - y) < s & r \\<le> abs (f z + -f y)) \
\ ==> \\<exists>X Y. \\<forall>n::nat. \
\ abs(X n + -(Y n)) < inverse(real(Suc n)) & \
\ r \\<le> abs(f (X n) + -f (Y n))";
@@ -745,23 +745,23 @@
Derivatives
------------------------------------------------------------------*)
Goalw [deriv_def]
- "(DERIV f x :> D) = ((%h. (f(x + h) + - f(x))/h) -- #0 --> D)";
+ "(DERIV f x :> D) = ((%h. (f(x + h) + - f(x))/h) -- Numeral0 --> D)";
by (Blast_tac 1);
qed "DERIV_iff";
Goalw [deriv_def]
- "(DERIV f x :> D) = ((%h. (f(x + h) + - f(x))/h) -- #0 --NS> D)";
+ "(DERIV f x :> D) = ((%h. (f(x + h) + - f(x))/h) -- Numeral0 --NS> D)";
by (simp_tac (simpset() addsimps [LIM_NSLIM_iff]) 1);
qed "DERIV_NS_iff";
Goalw [deriv_def]
"DERIV f x :> D \
-\ ==> (%h. (f(x + h) + - f(x))/h) -- #0 --> D";
+\ ==> (%h. (f(x + h) + - f(x))/h) -- Numeral0 --> D";
by (Blast_tac 1);
qed "DERIVD";
Goalw [deriv_def] "DERIV f x :> D ==> \
-\ (%h. (f(x + h) + - f(x))/h) -- #0 --NS> D";
+\ (%h. (f(x + h) + - f(x))/h) -- Numeral0 --NS> D";
by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff]) 1);
qed "NS_DERIVD";
@@ -809,7 +809,7 @@
-------------------------------------------------------*)
Goalw [LIM_def]
- "((%h. (f(a + h) + - f(a))/h) -- #0 --> D) = \
+ "((%h. (f(a + h) + - f(a))/h) -- Numeral0 --> D) = \
\ ((%x. (f(x) + -f(a)) / (x + -a)) -- a --> D)";
by Safe_tac;
by (ALLGOALS(dtac spec));
@@ -836,7 +836,7 @@
(*--- first equivalence ---*)
Goalw [nsderiv_def,NSLIM_def]
- "(NSDERIV f x :> D) = ((%h. (f(x + h) + - f(x))/h) -- #0 --NS> D)";
+ "(NSDERIV f x :> D) = ((%h. (f(x + h) + - f(x))/h) -- Numeral0 --NS> D)";
by (auto_tac (claset(), simpset() addsimps [hypreal_of_real_zero]));
by (dres_inst_tac [("x","xa")] bspec 1);
by (rtac ccontr 3);
@@ -956,12 +956,12 @@
------------------------*)
(* use simple constant nslimit theorem *)
-Goal "(NSDERIV (%x. k) x :> #0)";
+Goal "(NSDERIV (%x. k) x :> Numeral0)";
by (simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff]) 1);
qed "NSDERIV_const";
Addsimps [NSDERIV_const];
-Goal "(DERIV (%x. k) x :> #0)";
+Goal "(DERIV (%x. k) x :> Numeral0)";
by (simp_tac (simpset() addsimps [NSDERIV_DERIV_iff RS sym]) 1);
qed "DERIV_const";
Addsimps [DERIV_const];
@@ -1000,7 +1000,7 @@
Goal "[| (x + y) / z = hypreal_of_real D + yb; z \\<noteq> 0; \
\ z \\<in> Infinitesimal; yb \\<in> Infinitesimal |] \
-\ ==> x + y \\<approx> #0";
+\ ==> x + y \\<approx> Numeral0";
by (forw_inst_tac [("c1","z")] (hypreal_mult_right_cancel RS iffD2) 1
THEN assume_tac 1);
by (thin_tac "(x + y) / z = hypreal_of_real D + yb" 1);
@@ -1127,7 +1127,7 @@
qed "incrementI2";
(* The Increment theorem -- Keisler p. 65 *)
-Goal "[| NSDERIV f x :> D; h \\<in> Infinitesimal; h \\<noteq> #0 |] \
+Goal "[| NSDERIV f x :> D; h \\<in> Infinitesimal; h \\<noteq> Numeral0 |] \
\ ==> \\<exists>e \\<in> Infinitesimal. increment f x h = hypreal_of_real(D)*h + e*h";
by (forw_inst_tac [("h","h")] incrementI2 1 THEN rewtac nsderiv_def);
by (dtac bspec 1 THEN Auto_tac);
@@ -1143,15 +1143,15 @@
simpset() addsimps [hypreal_add_mult_distrib]));
qed "increment_thm";
-Goal "[| NSDERIV f x :> D; h \\<approx> #0; h \\<noteq> #0 |] \
+Goal "[| NSDERIV f x :> D; h \\<approx> Numeral0; h \\<noteq> Numeral0 |] \
\ ==> \\<exists>e \\<in> Infinitesimal. increment f x h = \
\ hypreal_of_real(D)*h + e*h";
by (blast_tac (claset() addSDs [mem_infmal_iff RS iffD2]
addSIs [increment_thm]) 1);
qed "increment_thm2";
-Goal "[| NSDERIV f x :> D; h \\<approx> #0; h \\<noteq> #0 |] \
-\ ==> increment f x h \\<approx> #0";
+Goal "[| NSDERIV f x :> D; h \\<approx> Numeral0; h \\<noteq> Numeral0 |] \
+\ ==> increment f x h \\<approx> Numeral0";
by (dtac increment_thm2 1 THEN auto_tac (claset() addSIs
[Infinitesimal_HFinite_mult2,HFinite_add],simpset() addsimps
[hypreal_add_mult_distrib RS sym,mem_infmal_iff RS sym]));
@@ -1172,16 +1172,16 @@
"[| NSDERIV g x :> D; \
\ (*f* g) (hypreal_of_real(x) + xa) = hypreal_of_real(g x);\
\ xa \\<in> Infinitesimal;\
-\ xa \\<noteq> #0 \
-\ |] ==> D = #0";
+\ xa \\<noteq> Numeral0 \
+\ |] ==> D = Numeral0";
by (dtac bspec 1);
by Auto_tac;
qed "NSDERIV_zero";
(* can be proved differently using NSLIM_isCont_iff *)
Goalw [nsderiv_def]
- "[| NSDERIV f x :> D; h \\<in> Infinitesimal; h \\<noteq> #0 |] \
-\ ==> (*f* f) (hypreal_of_real(x) + h) + -hypreal_of_real(f x) \\<approx> #0";
+ "[| NSDERIV f x :> D; h \\<in> Infinitesimal; h \\<noteq> Numeral0 |] \
+\ ==> (*f* f) (hypreal_of_real(x) + h) + -hypreal_of_real(f x) \\<approx> Numeral0";
by (asm_full_simp_tac (simpset() addsimps
[mem_infmal_iff RS sym]) 1);
by (rtac Infinitesimal_ratio 1);
@@ -1214,7 +1214,7 @@
----------------- \\<approx> Db
h
--------------------------------------------------------------*)
-Goal "[| NSDERIV g x :> Db; xa \\<in> Infinitesimal; xa \\<noteq> #0 |] \
+Goal "[| NSDERIV g x :> Db; xa \\<in> Infinitesimal; xa \\<noteq> Numeral0 |] \
\ ==> ((*f* g) (hypreal_of_real(x) + xa) + - hypreal_of_real(g x)) / xa \
\ \\<approx> hypreal_of_real(Db)";
by (auto_tac (claset(),
@@ -1266,7 +1266,7 @@
(*------------------------------------------------------------------
Differentiation of natural number powers
------------------------------------------------------------------*)
-Goal "NSDERIV (%x. x) x :> #1";
+Goal "NSDERIV (%x. x) x :> Numeral1";
by (auto_tac (claset(),
simpset() addsimps [NSDERIV_NSLIM_iff,
NSLIM_def ,starfun_Id, hypreal_of_real_zero,
@@ -1275,7 +1275,7 @@
Addsimps [NSDERIV_Id];
(*derivative of the identity function*)
-Goal "DERIV (%x. x) x :> #1";
+Goal "DERIV (%x. x) x :> Numeral1";
by (simp_tac (simpset() addsimps [NSDERIV_DERIV_iff RS sym]) 1);
qed "DERIV_Id";
Addsimps [DERIV_Id];
@@ -1294,7 +1294,7 @@
qed "NSDERIV_cmult_Id";
Addsimps [NSDERIV_cmult_Id];
-Goal "DERIV (%x. x ^ n) x :> real n * (x ^ (n - 1'))";
+Goal "DERIV (%x. x ^ n) x :> real n * (x ^ (n - Suc 0))";
by (induct_tac "n" 1);
by (dtac (DERIV_Id RS DERIV_mult) 2);
by (auto_tac (claset(),
@@ -1306,7 +1306,7 @@
qed "DERIV_pow";
(* NS version *)
-Goal "NSDERIV (%x. x ^ n) x :> real n * (x ^ (n - 1'))";
+Goal "NSDERIV (%x. x ^ n) x :> real n * (x ^ (n - Suc 0))";
by (simp_tac (simpset() addsimps [NSDERIV_DERIV_iff, DERIV_pow]) 1);
qed "NSDERIV_pow";
@@ -1314,9 +1314,9 @@
Power of -1
---------------------------------------------------------------*)
-(*Can't get rid of x \\<noteq> #0 because it isn't continuous at zero*)
+(*Can't get rid of x \\<noteq> Numeral0 because it isn't continuous at zero*)
Goalw [nsderiv_def]
- "x \\<noteq> #0 ==> NSDERIV (%x. inverse(x)) x :> (- (inverse x ^ 2))";
+ "x \\<noteq> Numeral0 ==> NSDERIV (%x. inverse(x)) x :> (- (inverse x ^ Suc (Suc 0)))";
by (rtac ballI 1 THEN Asm_full_simp_tac 1 THEN Step_tac 1);
by (forward_tac [Infinitesimal_add_not_zero] 1);
by (asm_full_simp_tac (simpset() addsimps [hypreal_add_commute]) 2);
@@ -1345,7 +1345,7 @@
qed "NSDERIV_inverse";
-Goal "x \\<noteq> #0 ==> DERIV (%x. inverse(x)) x :> (-(inverse x ^ 2))";
+Goal "x \\<noteq> Numeral0 ==> DERIV (%x. inverse(x)) x :> (-(inverse x ^ Suc (Suc 0)))";
by (asm_simp_tac (simpset() addsimps [NSDERIV_inverse,
NSDERIV_DERIV_iff RS sym] delsimps [realpow_Suc]) 1);
qed "DERIV_inverse";
@@ -1353,8 +1353,8 @@
(*--------------------------------------------------------------
Derivative of inverse
-------------------------------------------------------------*)
-Goal "[| DERIV f x :> d; f(x) \\<noteq> #0 |] \
-\ ==> DERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ 2)))";
+Goal "[| DERIV f x :> d; f(x) \\<noteq> Numeral0 |] \
+\ ==> DERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ Suc (Suc 0))))";
by (rtac (real_mult_commute RS subst) 1);
by (asm_simp_tac (simpset() addsimps [real_minus_mult_eq1,
realpow_inverse] delsimps [realpow_Suc,
@@ -1363,8 +1363,8 @@
by (blast_tac (claset() addSIs [DERIV_chain,DERIV_inverse]) 1);
qed "DERIV_inverse_fun";
-Goal "[| NSDERIV f x :> d; f(x) \\<noteq> #0 |] \
-\ ==> NSDERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ 2)))";
+Goal "[| NSDERIV f x :> d; f(x) \\<noteq> Numeral0 |] \
+\ ==> NSDERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ Suc (Suc 0))))";
by (asm_full_simp_tac (simpset() addsimps [NSDERIV_DERIV_iff,
DERIV_inverse_fun] delsimps [realpow_Suc]) 1);
qed "NSDERIV_inverse_fun";
@@ -1372,8 +1372,8 @@
(*--------------------------------------------------------------
Derivative of quotient
-------------------------------------------------------------*)
-Goal "[| DERIV f x :> d; DERIV g x :> e; g(x) \\<noteq> #0 |] \
-\ ==> DERIV (%y. f(y) / (g y)) x :> (d*g(x) + -(e*f(x))) / (g(x) ^ 2)";
+Goal "[| DERIV f x :> d; DERIV g x :> e; g(x) \\<noteq> Numeral0 |] \
+\ ==> DERIV (%y. f(y) / (g y)) x :> (d*g(x) + -(e*f(x))) / (g(x) ^ Suc (Suc 0))";
by (dres_inst_tac [("f","g")] DERIV_inverse_fun 1);
by (dtac DERIV_mult 2);
by (REPEAT(assume_tac 1));
@@ -1384,9 +1384,9 @@
real_minus_mult_eq2 RS sym]) 1);
qed "DERIV_quotient";
-Goal "[| NSDERIV f x :> d; DERIV g x :> e; g(x) \\<noteq> #0 |] \
+Goal "[| NSDERIV f x :> d; DERIV g x :> e; g(x) \\<noteq> Numeral0 |] \
\ ==> NSDERIV (%y. f(y) / (g y)) x :> (d*g(x) \
-\ + -(e*f(x))) / (g(x) ^ 2)";
+\ + -(e*f(x))) / (g(x) ^ Suc (Suc 0))";
by (asm_full_simp_tac (simpset() addsimps [NSDERIV_DERIV_iff,
DERIV_quotient] delsimps [realpow_Suc]) 1);
qed "NSDERIV_quotient";
@@ -1401,7 +1401,7 @@
by (res_inst_tac
[("x","%z. if z = x then l else (f(z) - f(x)) / (z - x)")] exI 1);
by (auto_tac (claset(),simpset() addsimps [real_mult_assoc,
- ARITH_PROVE "z \\<noteq> x ==> z - x \\<noteq> (#0::real)"]));
+ ARITH_PROVE "z \\<noteq> x ==> z - x \\<noteq> (Numeral0::real)"]));
by (auto_tac (claset(),simpset() addsimps [isCont_iff,DERIV_iff]));
by (ALLGOALS(rtac (LIM_equal RS iffD1)));
by (auto_tac (claset(),simpset() addsimps [real_diff_def,real_mult_assoc]));
@@ -1511,7 +1511,7 @@
Goal "[| \\<forall>n. f(n) \\<le> f(Suc n); \
\ \\<forall>n. g(Suc n) \\<le> g(n); \
\ \\<forall>n. f(n) \\<le> g(n); \
-\ (%n. f(n) - g(n)) ----> #0 |] \
+\ (%n. f(n) - g(n)) ----> Numeral0 |] \
\ ==> \\<exists>l. ((\\<forall>n. f(n) \\<le> l) & f ----> l) & \
\ ((\\<forall>n. l \\<le> g(n)) & g ----> l)";
by (dtac lemma_nest 1 THEN Auto_tac);
@@ -1544,15 +1544,15 @@
simpset() addsimps [Bolzano_bisect_le, Let_def, split_def]));
qed "Bolzano_bisect_Suc_le_snd";
-Goal "((x::real) = y / (#2 * z)) = (#2 * x = y/z)";
+Goal "((x::real) = y / (# 2 * z)) = (# 2 * x = y/z)";
by Auto_tac;
-by (dres_inst_tac [("f","%u. (#1/#2)*u")] arg_cong 1);
+by (dres_inst_tac [("f","%u. (Numeral1/# 2)*u")] arg_cong 1);
by Auto_tac;
qed "eq_divide_2_times_iff";
Goal "a \\<le> b ==> \
\ snd(Bolzano_bisect P a b n) - fst(Bolzano_bisect P a b n) = \
-\ (b-a) / (#2 ^ n)";
+\ (b-a) / (# 2 ^ n)";
by (induct_tac "n" 1);
by (auto_tac (claset(),
simpset() addsimps [eq_divide_2_times_iff, real_add_divide_distrib,
@@ -1589,7 +1589,7 @@
Goal "[| \\<forall>a b c. P(a,b) & P(b,c) & a \\<le> b & b \\<le> c --> P(a,c); \
-\ \\<forall>x. \\<exists>d::real. #0 < d & \
+\ \\<forall>x. \\<exists>d::real. Numeral0 < d & \
\ (\\<forall>a b. a \\<le> x & x \\<le> b & (b - a) < d --> P(a,b)); \
\ a \\<le> b |] \
\ ==> P(a,b)";
@@ -1604,8 +1604,8 @@
by (rename_tac "l" 1);
by (dres_inst_tac [("x","l")] spec 1 THEN Clarify_tac 1);
by (rewtac LIMSEQ_def);
-by (dres_inst_tac [("P", "%r. #0<r --> ?Q r"), ("x","d/#2")] spec 1);
-by (dres_inst_tac [("P", "%r. #0<r --> ?Q r"), ("x","d/#2")] spec 1);
+by (dres_inst_tac [("P", "%r. Numeral0<r --> ?Q r"), ("x","d/# 2")] spec 1);
+by (dres_inst_tac [("P", "%r. Numeral0<r --> ?Q r"), ("x","d/# 2")] spec 1);
by (dtac real_less_half_sum 1);
by Safe_tac;
(*linear arithmetic bug if we just use Asm_simp_tac*)
@@ -1626,7 +1626,7 @@
Goal "((\\<forall>a b c. (a \\<le> b & b \\<le> c & P(a,b) & P(b,c)) --> P(a,c)) & \
-\ (\\<forall>x. \\<exists>d::real. #0 < d & \
+\ (\\<forall>x. \\<exists>d::real. Numeral0 < d & \
\ (\\<forall>a b. a \\<le> x & x \\<le> b & (b - a) < d --> P(a,b)))) \
\ --> (\\<forall>a b. a \\<le> b --> P(a,b))";
by (Clarify_tac 1);
@@ -1654,14 +1654,14 @@
by (rtac ccontr 1);
by (subgoal_tac "a \\<le> x & x \\<le> b" 1);
by (Asm_full_simp_tac 2);
-by (dres_inst_tac [("P", "%d. #0<d --> ?P d"),("x","#1")] spec 2);
+by (dres_inst_tac [("P", "%d. Numeral0<d --> ?P d"),("x","Numeral1")] spec 2);
by (Step_tac 2);
by (Asm_full_simp_tac 2);
by (Asm_full_simp_tac 2);
by (REPEAT(blast_tac (claset() addIs [order_trans]) 2));
by (REPEAT(dres_inst_tac [("x","x")] spec 1));
by (Asm_full_simp_tac 1);
-by (dres_inst_tac [("P", "%r. ?P r --> (\\<exists>s. #0<s & ?Q r s)"),
+by (dres_inst_tac [("P", "%r. ?P r --> (\\<exists>s. Numeral0<s & ?Q r s)"),
("x","abs(y - f x)")] spec 1);
by Safe_tac;
by (asm_full_simp_tac (simpset() addsimps []) 1);
@@ -1738,15 +1738,15 @@
by (cut_inst_tac [("x","xb"),("y","xa")] linorder_linear 1);
by (Force_tac 1);
by (case_tac "a \\<le> x & x \\<le> b" 1);
-by (res_inst_tac [("x","#1")] exI 2);
+by (res_inst_tac [("x","Numeral1")] exI 2);
by (Force_tac 2);
by (asm_full_simp_tac (simpset() addsimps [LIM_def,isCont_iff]) 1);
by (dres_inst_tac [("x","x")] spec 1 THEN Auto_tac);
by (thin_tac "\\<forall>M. \\<exists>x. a \\<le> x & x \\<le> b & ~ f x \\<le> M" 1);
-by (dres_inst_tac [("x","#1")] spec 1);
+by (dres_inst_tac [("x","Numeral1")] spec 1);
by Auto_tac;
by (res_inst_tac [("x","s")] exI 1 THEN Clarify_tac 1);
-by (res_inst_tac [("x","abs(f x) + #1")] exI 1 THEN Clarify_tac 1);
+by (res_inst_tac [("x","abs(f x) + Numeral1")] exI 1 THEN Clarify_tac 1);
by (dres_inst_tac [("x","xa - x")] spec 1 THEN Safe_tac);
by (arith_tac 1);
by (arith_tac 1);
@@ -1803,23 +1803,23 @@
"\\<exists>k. \\<forall>x. a \\<le> x & x \\<le> b --> (%x. inverse(M - (f x))) x \\<le> k" 1);
by (rtac isCont_bounded 2);
by Safe_tac;
-by (subgoal_tac "\\<forall>x. a \\<le> x & x \\<le> b --> #0 < inverse(M - f(x))" 1);
+by (subgoal_tac "\\<forall>x. a \\<le> x & x \\<le> b --> Numeral0 < inverse(M - f(x))" 1);
by (Asm_full_simp_tac 1);
by Safe_tac;
by (asm_full_simp_tac (simpset() addsimps [real_less_diff_eq]) 2);
by (subgoal_tac
- "\\<forall>x. a \\<le> x & x \\<le> b --> (%x. inverse(M - (f x))) x < (k + #1)" 1);
+ "\\<forall>x. a \\<le> x & x \\<le> b --> (%x. inverse(M - (f x))) x < (k + Numeral1)" 1);
by Safe_tac;
by (res_inst_tac [("y","k")] order_le_less_trans 2);
by (asm_full_simp_tac (simpset() addsimps [real_zero_less_one]) 3);
by (Asm_full_simp_tac 2);
by (subgoal_tac "\\<forall>x. a \\<le> x & x \\<le> b --> \
-\ inverse(k + #1) < inverse((%x. inverse(M - (f x))) x)" 1);
+\ inverse(k + Numeral1) < inverse((%x. inverse(M - (f x))) x)" 1);
by Safe_tac;
by (rtac real_inverse_less_swap 2);
by (ALLGOALS Asm_full_simp_tac);
by (dres_inst_tac [("P", "%N. N<M --> ?Q N"),
- ("x","M - inverse(k + #1)")] spec 1);
+ ("x","M - inverse(k + Numeral1)")] spec 1);
by (Step_tac 1 THEN dtac real_leI 1);
by (dtac (real_le_diff_eq RS iffD1) 1);
by (REPEAT(dres_inst_tac [("x","a")] spec 1));
@@ -1879,11 +1879,11 @@
(*----------------------------------------------------------------------------*)
Goalw [deriv_def,LIM_def]
- "[| DERIV f x :> l; #0 < l |] ==> \
-\ \\<exists>d. #0 < d & (\\<forall>h. #0 < h & h < d --> f(x) < f(x + h))";
+ "[| DERIV f x :> l; Numeral0 < l |] ==> \
+\ \\<exists>d. Numeral0 < d & (\\<forall>h. Numeral0 < h & h < d --> f(x) < f(x + h))";
by (dtac spec 1 THEN Auto_tac);
by (res_inst_tac [("x","s")] exI 1 THEN Auto_tac);
-by (subgoal_tac "#0 < l*h" 1);
+by (subgoal_tac "Numeral0 < l*h" 1);
by (asm_full_simp_tac (simpset() addsimps [real_0_less_mult_iff]) 2);
by (dres_inst_tac [("x","h")] spec 1);
by (asm_full_simp_tac
@@ -1893,11 +1893,11 @@
qed "DERIV_left_inc";
Goalw [deriv_def,LIM_def]
- "[| DERIV f x :> l; l < #0 |] ==> \
-\ \\<exists>d. #0 < d & (\\<forall>h. #0 < h & h < d --> f(x) < f(x - h))";
+ "[| DERIV f x :> l; l < Numeral0 |] ==> \
+\ \\<exists>d. Numeral0 < d & (\\<forall>h. Numeral0 < h & h < d --> f(x) < f(x - h))";
by (dres_inst_tac [("x","-l")] spec 1 THEN Auto_tac);
by (res_inst_tac [("x","s")] exI 1 THEN Auto_tac);
-by (subgoal_tac "l*h < #0" 1);
+by (subgoal_tac "l*h < Numeral0" 1);
by (asm_full_simp_tac (simpset() addsimps [real_mult_less_0_iff]) 2);
by (dres_inst_tac [("x","-h")] spec 1);
by (asm_full_simp_tac
@@ -1905,7 +1905,7 @@
pos_real_less_divide_eq,
symmetric real_diff_def]
addsplits [split_if_asm]) 1);
-by (subgoal_tac "#0 < (f (x - h) - f x)/h" 1);
+by (subgoal_tac "Numeral0 < (f (x - h) - f x)/h" 1);
by (arith_tac 2);
by (asm_full_simp_tac
(simpset() addsimps [pos_real_less_divide_eq]) 1);
@@ -1913,9 +1913,9 @@
Goal "[| DERIV f x :> l; \
-\ \\<exists>d. #0 < d & (\\<forall>y. abs(x - y) < d --> f(y) \\<le> f(x)) |] \
-\ ==> l = #0";
-by (res_inst_tac [("R1.0","l"),("R2.0","#0")] real_linear_less2 1);
+\ \\<exists>d. Numeral0 < d & (\\<forall>y. abs(x - y) < d --> f(y) \\<le> f(x)) |] \
+\ ==> l = Numeral0";
+by (res_inst_tac [("R1.0","l"),("R2.0","Numeral0")] real_linear_less2 1);
by Safe_tac;
by (dtac DERIV_left_dec 1);
by (dtac DERIV_left_inc 3);
@@ -1933,8 +1933,8 @@
(*----------------------------------------------------------------------------*)
Goal "[| DERIV f x :> l; \
-\ \\<exists>d::real. #0 < d & (\\<forall>y. abs(x - y) < d --> f(x) \\<le> f(y)) |] \
-\ ==> l = #0";
+\ \\<exists>d::real. Numeral0 < d & (\\<forall>y. abs(x - y) < d --> f(x) \\<le> f(y)) |] \
+\ ==> l = Numeral0";
by (dtac (DERIV_minus RS DERIV_local_max) 1);
by Auto_tac;
qed "DERIV_local_min";
@@ -1944,8 +1944,8 @@
(*----------------------------------------------------------------------------*)
Goal "[| DERIV f x :> l; \
-\ \\<exists>d. #0 < d & (\\<forall>y. abs(x - y) < d --> f(x) = f(y)) |] \
-\ ==> l = #0";
+\ \\<exists>d. Numeral0 < d & (\\<forall>y. abs(x - y) < d --> f(x) = f(y)) |] \
+\ ==> l = Numeral0";
by (auto_tac (claset() addSDs [DERIV_local_max],simpset()));
qed "DERIV_local_const";
@@ -1954,7 +1954,7 @@
(*----------------------------------------------------------------------------*)
Goal "[| a < x; x < b |] ==> \
-\ \\<exists>d::real. #0 < d & (\\<forall>y. abs(x - y) < d --> a < y & y < b)";
+\ \\<exists>d::real. Numeral0 < d & (\\<forall>y. abs(x - y) < d --> a < y & y < b)";
by (simp_tac (simpset() addsimps [abs_interval_iff]) 1);
by (cut_inst_tac [("x","x - a"),("y","b - x")] linorder_linear 1);
by Safe_tac;
@@ -1965,7 +1965,7 @@
qed "lemma_interval_lt";
Goal "[| a < x; x < b |] ==> \
-\ \\<exists>d::real. #0 < d & (\\<forall>y. abs(x - y) < d --> a \\<le> y & y \\<le> b)";
+\ \\<exists>d::real. Numeral0 < d & (\\<forall>y. abs(x - y) < d --> a \\<le> y & y \\<le> b)";
by (dtac lemma_interval_lt 1);
by Auto_tac;
by (auto_tac (claset() addSIs [exI] ,simpset()));
@@ -1975,13 +1975,13 @@
Rolle's Theorem
If f is defined and continuous on the finite closed interval [a,b]
and differentiable a least on the open interval (a,b), and f(a) = f(b),
- then x0 \\<in> (a,b) such that f'(x0) = #0
+ then x0 \\<in> (a,b) such that f'(x0) = Numeral0
----------------------------------------------------------------------*)
Goal "[| a < b; f(a) = f(b); \
\ \\<forall>x. a \\<le> x & x \\<le> b --> isCont f x; \
\ \\<forall>x. a < x & x < b --> f differentiable x \
-\ |] ==> \\<exists>z. a < z & z < b & DERIV f z :> #0";
+\ |] ==> \\<exists>z. a < z & z < b & DERIV f z :> Numeral0";
by (ftac (order_less_imp_le RS isCont_eq_Ub) 1);
by (EVERY1[assume_tac,Step_tac]);
by (ftac (order_less_imp_le RS isCont_eq_Lb) 1);
@@ -1992,7 +1992,7 @@
by (EVERY1[assume_tac,etac exE]);
by (res_inst_tac [("x","x")] exI 1 THEN Asm_full_simp_tac 1);
by (subgoal_tac "(\\<exists>l. DERIV f x :> l) & \
-\ (\\<exists>d. #0 < d & (\\<forall>y. abs(x - y) < d --> f(y) \\<le> f(x)))" 1);
+\ (\\<exists>d. Numeral0 < d & (\\<forall>y. abs(x - y) < d --> f(y) \\<le> f(x)))" 1);
by (Clarify_tac 1 THEN rtac conjI 2);
by (blast_tac (claset() addIs [differentiableD]) 2);
by (Blast_tac 2);
@@ -2004,7 +2004,7 @@
by (EVERY1[assume_tac,etac exE]);
by (res_inst_tac [("x","xa")] exI 1 THEN Asm_full_simp_tac 1);
by (subgoal_tac "(\\<exists>l. DERIV f xa :> l) & \
-\ (\\<exists>d. #0 < d & (\\<forall>y. abs(xa - y) < d --> f(xa) \\<le> f(y)))" 1);
+\ (\\<exists>d. Numeral0 < d & (\\<forall>y. abs(xa - y) < d --> f(xa) \\<le> f(y)))" 1);
by (Clarify_tac 1 THEN rtac conjI 2);
by (blast_tac (claset() addIs [differentiableD]) 2);
by (Blast_tac 2);
@@ -2030,7 +2030,7 @@
by (forw_inst_tac [("a","a"),("x","r")] lemma_interval 1);
by (EVERY1[assume_tac, etac exE]);
by (subgoal_tac "(\\<exists>l. DERIV f r :> l) & \
-\ (\\<exists>d. #0 < d & (\\<forall>y. abs(r - y) < d --> f(r) = f(y)))" 1);
+\ (\\<exists>d. Numeral0 < d & (\\<forall>y. abs(r - y) < d --> f(r) = f(y)))" 1);
by (Clarify_tac 1 THEN rtac conjI 2);
by (blast_tac (claset() addIs [differentiableD]) 2);
by (EVERY1[ftac DERIV_local_const, Blast_tac, Blast_tac]);
@@ -2098,7 +2098,7 @@
Goal "[| a < b; \
\ \\<forall>x. a \\<le> x & x \\<le> b --> isCont f x; \
-\ \\<forall>x. a < x & x < b --> DERIV f x :> #0 |] \
+\ \\<forall>x. a < x & x < b --> DERIV f x :> Numeral0 |] \
\ ==> (f b = f a)";
by (dtac MVT 1 THEN assume_tac 1);
by (blast_tac (claset() addIs [differentiableI]) 1);
@@ -2108,7 +2108,7 @@
Goal "[| a < b; \
\ \\<forall>x. a \\<le> x & x \\<le> b --> isCont f x; \
-\ \\<forall>x. a < x & x < b --> DERIV f x :> #0 |] \
+\ \\<forall>x. a < x & x < b --> DERIV f x :> Numeral0 |] \
\ ==> \\<forall>x. a \\<le> x & x \\<le> b --> f x = f a";
by Safe_tac;
by (dres_inst_tac [("x","a")] order_le_imp_less_or_eq 1);
@@ -2119,13 +2119,13 @@
Goal "[| a < b; \
\ \\<forall>x. a \\<le> x & x \\<le> b --> isCont f x; \
-\ \\<forall>x. a < x & x < b --> DERIV f x :> #0; \
+\ \\<forall>x. a < x & x < b --> DERIV f x :> Numeral0; \
\ a \\<le> x; x \\<le> b |] \
\ ==> f x = f a";
by (blast_tac (claset() addDs [DERIV_isconst1]) 1);
qed "DERIV_isconst2";
-Goal "\\<forall>x. DERIV f x :> #0 ==> f(x) = f(y)";
+Goal "\\<forall>x. DERIV f x :> Numeral0 ==> f(x) = f(y)";
by (res_inst_tac [("R1.0","x"),("R2.0","y")] real_linear_less2 1);
by (rtac sym 1);
by (auto_tac (claset() addIs [DERIV_isCont,DERIV_isconst_end],simpset()));
@@ -2148,12 +2148,12 @@
simpset() addsimps [real_mult_assoc]));
qed "DERIV_const_ratio_const2";
-Goal "((a + b) /#2 - a) = (b - a)/(#2::real)";
+Goal "((a + b) /# 2 - a) = (b - a)/(# 2::real)";
by Auto_tac;
qed "real_average_minus_first";
Addsimps [real_average_minus_first];
-Goal "((b + a)/#2 - a) = (b - a)/(#2::real)";
+Goal "((b + a)/# 2 - a) = (b - a)/(# 2::real)";
by Auto_tac;
qed "real_average_minus_second";
Addsimps [real_average_minus_second];
@@ -2161,7 +2161,7 @@
(* Gallileo's "trick": average velocity = av. of end velocities *)
Goal "[|a \\<noteq> (b::real); \\<forall>x. DERIV v x :> k|] \
-\ ==> v((a + b)/#2) = (v a + v b)/#2";
+\ ==> v((a + b)/# 2) = (v a + v b)/# 2";
by (res_inst_tac [("R1.0","a"),("R2.0","b")] real_linear_less2 1);
by Auto_tac;
by (ftac DERIV_const_ratio_const2 1 THEN assume_tac 1);
@@ -2182,7 +2182,7 @@
(* maximum at an end point, not in the middle. *)
(* ------------------------------------------------------------------------ *)
-Goal "[|#0 < d; \\<forall>z. abs(z - x) \\<le> d --> g(f z) = z; \
+Goal "[|Numeral0 < d; \\<forall>z. abs(z - x) \\<le> d --> g(f z) = z; \
\ \\<forall>z. abs(z - x) \\<le> d --> isCont f z |] \
\ ==> ~(\\<forall>z. abs(z - x) \\<le> d --> f(z) \\<le> f(x))";
by (rtac notI 1);
@@ -2221,7 +2221,7 @@
(* Similar version for lower bound *)
(* ------------------------------------------------------------------------ *)
-Goal "[|#0 < d; \\<forall>z. abs(z - x) \\<le> d --> g(f z) = z; \
+Goal "[|Numeral0 < d; \\<forall>z. abs(z - x) \\<le> d --> g(f z) = z; \
\ \\<forall>z. abs(z - x) \\<le> d --> isCont f z |] \
\ ==> ~(\\<forall>z. abs(z - x) \\<le> d --> f(x) \\<le> f(z))";
by (auto_tac (claset() addSDs [(asm_full_simplify (simpset())
@@ -2236,12 +2236,12 @@
Addsimps [zero_eq_numeral_0,one_eq_numeral_1];
-val lemma_le = ARITH_PROVE "#0 \\<le> (d::real) ==> -d \\<le> d";
+val lemma_le = ARITH_PROVE "Numeral0 \\<le> (d::real) ==> -d \\<le> d";
(* FIXME: awful proof - needs improvement *)
-Goal "[| #0 < d; \\<forall>z. abs(z - x) \\<le> d --> g(f z) = z; \
+Goal "[| Numeral0 < d; \\<forall>z. abs(z - x) \\<le> d --> g(f z) = z; \
\ \\<forall>z. abs(z - x) \\<le> d --> isCont f z |] \
-\ ==> \\<exists>e. #0 < e & \
+\ ==> \\<exists>e. Numeral0 < e & \
\ (\\<forall>y. \
\ abs(y - f(x)) \\<le> e --> \
\ (\\<exists>z. abs(z - x) \\<le> d & (f z = y)))";
@@ -2255,8 +2255,8 @@
by (Asm_full_simp_tac 2);
by (subgoal_tac "L < f x & f x < M" 1);
by Safe_tac;
-by (dres_inst_tac [("x","L")] (ARITH_PROVE "x < y ==> #0 < y - (x::real)") 1);
-by (dres_inst_tac [("x","f x")] (ARITH_PROVE "x < y ==> #0 < y - (x::real)") 1);
+by (dres_inst_tac [("x","L")] (ARITH_PROVE "x < y ==> Numeral0 < y - (x::real)") 1);
+by (dres_inst_tac [("x","f x")] (ARITH_PROVE "x < y ==> Numeral0 < y - (x::real)") 1);
by (dres_inst_tac [("d1.0","f x - L"),("d2.0","M - f x")]
(rename_numerals real_lbound_gt_zero) 1);
by Safe_tac;
@@ -2284,7 +2284,7 @@
(* Continuity of inverse function *)
(* ------------------------------------------------------------------------ *)
-Goal "[| #0 < d; \\<forall>z. abs(z - x) \\<le> d --> g(f(z)) = z; \
+Goal "[| Numeral0 < d; \\<forall>z. abs(z - x) \\<le> d --> g(f(z)) = z; \
\ \\<forall>z. abs(z - x) \\<le> d --> isCont f z |] \
\ ==> isCont g (f x)";
by (simp_tac (simpset() addsimps [isCont_iff,LIM_def]) 1);
--- a/src/HOL/Hyperreal/Lim.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Hyperreal/Lim.thy Fri Oct 05 21:52:39 2001 +0200
@@ -15,8 +15,8 @@
LIM :: [real=>real,real,real] => bool
("((_)/ -- (_)/ --> (_))" [60, 0, 60] 60)
"f -- a --> L ==
- ALL r. #0 < r -->
- (EX s. #0 < s & (ALL x. (x ~= a & (abs(x + -a) < s)
+ ALL r. Numeral0 < r -->
+ (EX s. Numeral0 < s & (ALL x. (x ~= a & (abs(x + -a) < s)
--> abs(f x + -L) < r)))"
NSLIM :: [real=>real,real,real] => bool
@@ -36,7 +36,7 @@
(* differentiation: D is derivative of function f at x *)
deriv:: [real=>real,real,real] => bool
("(DERIV (_)/ (_)/ :> (_))" [60, 0, 60] 60)
- "DERIV f x :> D == ((%h. (f(x + h) + -f(x))/h) -- #0 --> D)"
+ "DERIV f x :> D == ((%h. (f(x + h) + -f(x))/h) -- Numeral0 --> D)"
nsderiv :: [real=>real,real,real] => bool
("(NSDERIV (_)/ (_)/ :> (_))" [60, 0, 60] 60)
@@ -55,8 +55,8 @@
inc = (*f* f)(hypreal_of_real x + h) + -hypreal_of_real (f x))"
isUCont :: (real=>real) => bool
- "isUCont f == (ALL r. #0 < r -->
- (EX s. #0 < s & (ALL x y. abs(x + -y) < s
+ "isUCont f == (ALL r. Numeral0 < r -->
+ (EX s. Numeral0 < s & (ALL x y. abs(x + -y) < s
--> abs(f x + -f y) < r)))"
isNSUCont :: (real=>real) => bool
@@ -71,8 +71,8 @@
"Bolzano_bisect P a b 0 = (a,b)"
"Bolzano_bisect P a b (Suc n) =
(let (x,y) = Bolzano_bisect P a b n
- in if P(x, (x+y)/#2) then ((x+y)/#2, y)
- else (x, (x+y)/#2) )"
+ in if P(x, (x+y)/# 2) then ((x+y)/# 2, y)
+ else (x, (x+y)/# 2) )"
end
--- a/src/HOL/Hyperreal/NSA.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Hyperreal/NSA.ML Fri Oct 05 21:52:39 2001 +0200
@@ -211,9 +211,9 @@
Goalw [SReal_def,HFinite_def] "Reals <= HFinite";
by Auto_tac;
-by (res_inst_tac [("x","#1 + abs(hypreal_of_real r)")] exI 1);
+by (res_inst_tac [("x","Numeral1 + abs(hypreal_of_real r)")] exI 1);
by (auto_tac (claset(), simpset() addsimps [hypreal_of_real_hrabs]));
-by (res_inst_tac [("x","#1 + abs r")] exI 1);
+by (res_inst_tac [("x","Numeral1 + abs r")] exI 1);
by (Simp_tac 1);
qed "SReal_subset_HFinite";
@@ -238,8 +238,8 @@
qed "HFinite_number_of";
Addsimps [HFinite_number_of];
-Goal "[|x : HFinite; y <= x; #0 <= y |] ==> y: HFinite";
-by (case_tac "x <= #0" 1);
+Goal "[|x : HFinite; y <= x; Numeral0 <= y |] ==> y: HFinite";
+by (case_tac "x <= Numeral0" 1);
by (dres_inst_tac [("y","x")] order_trans 1);
by (dtac hypreal_le_anti_sym 2);
by (auto_tac (claset() addSDs [not_hypreal_leE], simpset()));
@@ -251,20 +251,20 @@
Set of infinitesimals is a subring of the hyperreals
------------------------------------------------------------------*)
Goalw [Infinitesimal_def]
- "x : Infinitesimal ==> ALL r: Reals. #0 < r --> abs x < r";
+ "x : Infinitesimal ==> ALL r: Reals. Numeral0 < r --> abs x < r";
by Auto_tac;
qed "InfinitesimalD";
-Goalw [Infinitesimal_def] "#0 : Infinitesimal";
+Goalw [Infinitesimal_def] "Numeral0 : Infinitesimal";
by (simp_tac (simpset() addsimps [hrabs_zero]) 1);
qed "Infinitesimal_zero";
AddIffs [Infinitesimal_zero];
-Goal "x/(#2::hypreal) + x/(#2::hypreal) = x";
+Goal "x/(# 2::hypreal) + x/(# 2::hypreal) = x";
by Auto_tac;
qed "hypreal_sum_of_halves";
-Goal "#0 < r ==> #0 < r/(#2::hypreal)";
+Goal "Numeral0 < r ==> Numeral0 < r/(# 2::hypreal)";
by Auto_tac;
qed "hypreal_half_gt_zero";
@@ -290,8 +290,8 @@
Goalw [Infinitesimal_def]
"[| x : Infinitesimal; y : Infinitesimal |] ==> (x * y) : Infinitesimal";
by Auto_tac;
-by (case_tac "y=#0" 1);
-by (cut_inst_tac [("u","abs x"),("v","#1"),("x","abs y"),("y","r")]
+by (case_tac "y=Numeral0" 1);
+by (cut_inst_tac [("u","abs x"),("v","Numeral1"),("x","abs y"),("y","r")]
hypreal_mult_less_mono 2);
by Auto_tac;
qed "Infinitesimal_mult";
@@ -332,27 +332,27 @@
Goalw [HInfinite_def] "[|x: HInfinite;y: HInfinite|] ==> (x*y): HInfinite";
by Auto_tac;
-by (eres_inst_tac [("x","#1")] ballE 1);
+by (eres_inst_tac [("x","Numeral1")] ballE 1);
by (eres_inst_tac [("x","r")] ballE 1);
by (case_tac "y=0" 1);
-by (cut_inst_tac [("x","#1"),("y","abs x"),
+by (cut_inst_tac [("x","Numeral1"),("y","abs x"),
("u","r"),("v","abs y")] hypreal_mult_less_mono 2);
by (auto_tac (claset(), simpset() addsimps hypreal_mult_ac));
qed "HInfinite_mult";
Goalw [HInfinite_def]
- "[|x: HInfinite; #0 <= y; #0 <= x|] ==> (x + y): HInfinite";
+ "[|x: HInfinite; Numeral0 <= y; Numeral0 <= x|] ==> (x + y): HInfinite";
by (auto_tac (claset() addSIs [hypreal_add_zero_less_le_mono],
simpset() addsimps [hrabs_eqI1, hypreal_add_commute,
hypreal_le_add_order]));
qed "HInfinite_add_ge_zero";
-Goal "[|x: HInfinite; #0 <= y; #0 <= x|] ==> (y + x): HInfinite";
+Goal "[|x: HInfinite; Numeral0 <= y; Numeral0 <= x|] ==> (y + x): HInfinite";
by (auto_tac (claset() addSIs [HInfinite_add_ge_zero],
simpset() addsimps [hypreal_add_commute]));
qed "HInfinite_add_ge_zero2";
-Goal "[|x: HInfinite; #0 < y; #0 < x|] ==> (x + y): HInfinite";
+Goal "[|x: HInfinite; Numeral0 < y; Numeral0 < x|] ==> (x + y): HInfinite";
by (blast_tac (claset() addIs [HInfinite_add_ge_zero,
order_less_imp_le]) 1);
qed "HInfinite_add_gt_zero";
@@ -361,14 +361,14 @@
by Auto_tac;
qed "HInfinite_minus_iff";
-Goal "[|x: HInfinite; y <= #0; x <= #0|] ==> (x + y): HInfinite";
+Goal "[|x: HInfinite; y <= Numeral0; x <= Numeral0|] ==> (x + y): HInfinite";
by (dtac (HInfinite_minus_iff RS iffD2) 1);
by (rtac (HInfinite_minus_iff RS iffD1) 1);
by (auto_tac (claset() addIs [HInfinite_add_ge_zero],
simpset() addsimps [hypreal_minus_zero_le_iff]));
qed "HInfinite_add_le_zero";
-Goal "[|x: HInfinite; y < #0; x < #0|] ==> (x + y): HInfinite";
+Goal "[|x: HInfinite; y < Numeral0; x < Numeral0|] ==> (x + y): HInfinite";
by (blast_tac (claset() addIs [HInfinite_add_le_zero,
order_less_imp_le]) 1);
qed "HInfinite_add_lt_zero";
@@ -378,11 +378,11 @@
by (auto_tac (claset() addIs [HFinite_mult,HFinite_add], simpset()));
qed "HFinite_sum_squares";
-Goal "x ~: Infinitesimal ==> x ~= #0";
+Goal "x ~: Infinitesimal ==> x ~= Numeral0";
by Auto_tac;
qed "not_Infinitesimal_not_zero";
-Goal "x: HFinite - Infinitesimal ==> x ~= #0";
+Goal "x: HFinite - Infinitesimal ==> x ~= Numeral0";
by Auto_tac;
qed "not_Infinitesimal_not_zero2";
@@ -441,7 +441,7 @@
by (fast_tac (claset() addDs [not_Infinitesimal_mult]) 1);
qed "Infinitesimal_mult_disj";
-Goal "x: HFinite-Infinitesimal ==> x ~= #0";
+Goal "x: HFinite-Infinitesimal ==> x ~= Numeral0";
by (Blast_tac 1);
qed "HFinite_Infinitesimal_not_zero";
@@ -455,7 +455,7 @@
Goalw [Infinitesimal_def,HFinite_def]
"Infinitesimal <= HFinite";
by Auto_tac;
-by (res_inst_tac [("x","#1")] bexI 1);
+by (res_inst_tac [("x","Numeral1")] bexI 1);
by Auto_tac;
qed "Infinitesimal_subset_HFinite";
@@ -474,15 +474,15 @@
----------------------------------------------------------------------*)
Goalw [Infinitesimal_def,approx_def]
- "(x:Infinitesimal) = (x @= #0)";
+ "(x:Infinitesimal) = (x @= Numeral0)";
by (Simp_tac 1);
qed "mem_infmal_iff";
-Goalw [approx_def]" (x @= y) = (x + -y @= #0)";
+Goalw [approx_def]" (x @= y) = (x + -y @= Numeral0)";
by (Simp_tac 1);
qed "approx_minus_iff";
-Goalw [approx_def]" (x @= y) = (-y + x @= #0)";
+Goalw [approx_def]" (x @= y) = (-y + x @= Numeral0)";
by (simp_tac (simpset() addsimps [hypreal_add_commute]) 1);
qed "approx_minus_iff2";
@@ -704,36 +704,36 @@
approx_hypreal_of_real_HFinite,HFinite_hypreal_of_real]) 1);
qed "approx_mult_hypreal_of_real";
-Goal "[| a: Reals; a ~= #0; a*x @= #0 |] ==> x @= #0";
+Goal "[| a: Reals; a ~= Numeral0; a*x @= Numeral0 |] ==> x @= Numeral0";
by (dtac (SReal_inverse RS (SReal_subset_HFinite RS subsetD)) 1);
by (auto_tac (claset() addDs [approx_mult2],
simpset() addsimps [hypreal_mult_assoc RS sym]));
qed "approx_SReal_mult_cancel_zero";
(* REM comments: newly added *)
-Goal "[| a: Reals; x @= #0 |] ==> x*a @= #0";
+Goal "[| a: Reals; x @= Numeral0 |] ==> x*a @= Numeral0";
by (auto_tac (claset() addDs [(SReal_subset_HFinite RS subsetD),
approx_mult1], simpset()));
qed "approx_mult_SReal1";
-Goal "[| a: Reals; x @= #0 |] ==> a*x @= #0";
+Goal "[| a: Reals; x @= Numeral0 |] ==> a*x @= Numeral0";
by (auto_tac (claset() addDs [(SReal_subset_HFinite RS subsetD),
approx_mult2], simpset()));
qed "approx_mult_SReal2";
-Goal "[|a : Reals; a ~= #0 |] ==> (a*x @= #0) = (x @= #0)";
+Goal "[|a : Reals; a ~= Numeral0 |] ==> (a*x @= Numeral0) = (x @= Numeral0)";
by (blast_tac (claset() addIs [approx_SReal_mult_cancel_zero,
approx_mult_SReal2]) 1);
qed "approx_mult_SReal_zero_cancel_iff";
Addsimps [approx_mult_SReal_zero_cancel_iff];
-Goal "[| a: Reals; a ~= #0; a* w @= a*z |] ==> w @= z";
+Goal "[| a: Reals; a ~= Numeral0; a* w @= a*z |] ==> w @= z";
by (dtac (SReal_inverse RS (SReal_subset_HFinite RS subsetD)) 1);
by (auto_tac (claset() addDs [approx_mult2],
simpset() addsimps [hypreal_mult_assoc RS sym]));
qed "approx_SReal_mult_cancel";
-Goal "[| a: Reals; a ~= #0|] ==> (a* w @= a*z) = (w @= z)";
+Goal "[| a: Reals; a ~= Numeral0|] ==> (a* w @= a*z) = (w @= z)";
by (auto_tac (claset() addSIs [approx_mult2,SReal_subset_HFinite RS subsetD]
addIs [approx_SReal_mult_cancel], simpset()));
qed "approx_SReal_mult_cancel_iff1";
@@ -754,50 +754,50 @@
-----------------------------------------------------------------*)
Goalw [Infinitesimal_def]
- "[| x: Reals; y: Infinitesimal; #0 < x |] ==> y < x";
+ "[| x: Reals; y: Infinitesimal; Numeral0 < x |] ==> y < x";
by (rtac (hrabs_ge_self RS order_le_less_trans) 1);
by Auto_tac;
qed "Infinitesimal_less_SReal";
-Goal "y: Infinitesimal ==> ALL r: Reals. #0 < r --> y < r";
+Goal "y: Infinitesimal ==> ALL r: Reals. Numeral0 < r --> y < r";
by (blast_tac (claset() addIs [Infinitesimal_less_SReal]) 1);
qed "Infinitesimal_less_SReal2";
Goalw [Infinitesimal_def]
- "[| #0 < y; y: Reals|] ==> y ~: Infinitesimal";
+ "[| Numeral0 < y; y: Reals|] ==> y ~: Infinitesimal";
by (auto_tac (claset(), simpset() addsimps [hrabs_def]));
qed "SReal_not_Infinitesimal";
-Goal "[| y < #0; y : Reals |] ==> y ~: Infinitesimal";
+Goal "[| y < Numeral0; y : Reals |] ==> y ~: Infinitesimal";
by (stac (Infinitesimal_minus_iff RS sym) 1);
by (rtac SReal_not_Infinitesimal 1);
by Auto_tac;
qed "SReal_minus_not_Infinitesimal";
-Goal "Reals Int Infinitesimal = {#0}";
+Goal "Reals Int Infinitesimal = {Numeral0}";
by Auto_tac;
-by (cut_inst_tac [("x","x"),("y","#0")] hypreal_linear 1);
+by (cut_inst_tac [("x","x"),("y","Numeral0")] hypreal_linear 1);
by (blast_tac (claset() addDs [SReal_not_Infinitesimal,
SReal_minus_not_Infinitesimal]) 1);
qed "SReal_Int_Infinitesimal_zero";
-Goal "[| x: Reals; x: Infinitesimal|] ==> x = #0";
+Goal "[| x: Reals; x: Infinitesimal|] ==> x = Numeral0";
by (cut_facts_tac [SReal_Int_Infinitesimal_zero] 1);
by (Blast_tac 1);
qed "SReal_Infinitesimal_zero";
-Goal "[| x : Reals; x ~= #0 |] ==> x : HFinite - Infinitesimal";
+Goal "[| x : Reals; x ~= Numeral0 |] ==> x : HFinite - Infinitesimal";
by (auto_tac (claset() addDs [SReal_Infinitesimal_zero,
SReal_subset_HFinite RS subsetD],
simpset()));
qed "SReal_HFinite_diff_Infinitesimal";
-Goal "hypreal_of_real x ~= #0 ==> hypreal_of_real x : HFinite - Infinitesimal";
+Goal "hypreal_of_real x ~= Numeral0 ==> hypreal_of_real x : HFinite - Infinitesimal";
by (rtac SReal_HFinite_diff_Infinitesimal 1);
by Auto_tac;
qed "hypreal_of_real_HFinite_diff_Infinitesimal";
-Goal "(hypreal_of_real x : Infinitesimal) = (x=#0)";
+Goal "(hypreal_of_real x : Infinitesimal) = (x=Numeral0)";
by (auto_tac (claset(), simpset() addsimps [hypreal_of_real_zero]));
by (rtac ccontr 1);
by (rtac (hypreal_of_real_HFinite_diff_Infinitesimal RS DiffD2) 1);
@@ -805,12 +805,12 @@
qed "hypreal_of_real_Infinitesimal_iff_0";
AddIffs [hypreal_of_real_Infinitesimal_iff_0];
-Goal "number_of w ~= (#0::hypreal) ==> number_of w ~: Infinitesimal";
+Goal "number_of w ~= (Numeral0::hypreal) ==> number_of w ~: Infinitesimal";
by (fast_tac (claset() addDs [SReal_number_of RS SReal_Infinitesimal_zero]) 1);
qed "number_of_not_Infinitesimal";
Addsimps [number_of_not_Infinitesimal];
-Goal "[| y: Reals; x @= y; y~= #0 |] ==> x ~= #0";
+Goal "[| y: Reals; x @= y; y~= Numeral0 |] ==> x ~= Numeral0";
by (cut_inst_tac [("x","y")] hypreal_trichotomy 1);
by (Asm_full_simp_tac 1);
by (blast_tac (claset() addDs
@@ -828,7 +828,7 @@
(*The premise y~=0 is essential; otherwise x/y =0 and we lose the
HFinite premise.*)
-Goal "[| y ~= #0; y: Infinitesimal; x/y : HFinite |] ==> x : Infinitesimal";
+Goal "[| y ~= Numeral0; y: Infinitesimal; x/y : HFinite |] ==> x : Infinitesimal";
by (dtac Infinitesimal_HFinite_mult2 1);
by (assume_tac 1);
by (asm_full_simp_tac
@@ -912,7 +912,7 @@
lemma_st_part_nonempty, lemma_st_part_subset]) 1);
qed "lemma_st_part_lub";
-Goal "((t::hypreal) + r <= t) = (r <= #0)";
+Goal "((t::hypreal) + r <= t) = (r <= Numeral0)";
by (Step_tac 1);
by (dres_inst_tac [("x","-t")] hypreal_add_left_le_mono1 1);
by (dres_inst_tac [("x","t")] hypreal_add_left_le_mono1 2);
@@ -920,7 +920,7 @@
qed "lemma_hypreal_le_left_cancel";
Goal "[| x: HFinite; isLub Reals {s. s: Reals & s < x} t; \
-\ r: Reals; #0 < r |] ==> x <= t + r";
+\ r: Reals; Numeral0 < r |] ==> x <= t + r";
by (forward_tac [isLubD1a] 1);
by (rtac ccontr 1 THEN dtac (linorder_not_le RS iffD2) 1);
by (dres_inst_tac [("x","t")] SReal_add 1 THEN assume_tac 1);
@@ -945,14 +945,14 @@
addIs [order_less_imp_le] addSIs [isUbI,setleI], simpset()));
qed "lemma_st_part_gt_ub";
-Goal "t <= t + -r ==> r <= (#0::hypreal)";
+Goal "t <= t + -r ==> r <= (Numeral0::hypreal)";
by (dres_inst_tac [("x","-t")] hypreal_add_left_le_mono1 1);
by (auto_tac (claset(), simpset() addsimps [hypreal_add_assoc RS sym]));
qed "lemma_minus_le_zero";
Goal "[| x: HFinite; \
\ isLub Reals {s. s: Reals & s < x} t; \
-\ r: Reals; #0 < r |] \
+\ r: Reals; Numeral0 < r |] \
\ ==> t + -r <= x";
by (forward_tac [isLubD1a] 1);
by (rtac ccontr 1 THEN dtac not_hypreal_leE 1);
@@ -970,7 +970,7 @@
Goal "[| x: HFinite; \
\ isLub Reals {s. s: Reals & s < x} t; \
-\ r: Reals; #0 < r |] \
+\ r: Reals; Numeral0 < r |] \
\ ==> x + -t <= r";
by (blast_tac (claset() addSIs [lemma_hypreal_le_swap RS iffD1,
lemma_st_part_le1]) 1);
@@ -982,7 +982,7 @@
Goal "[| x: HFinite; \
\ isLub Reals {s. s: Reals & s < x} t; \
-\ r: Reals; #0 < r |] \
+\ r: Reals; Numeral0 < r |] \
\ ==> -(x + -t) <= r";
by (blast_tac (claset() addSIs [lemma_hypreal_le_swap2 RS iffD1,
lemma_st_part_le2]) 1);
@@ -1004,7 +1004,7 @@
Goal "[| x: HFinite; \
\ isLub Reals {s. s: Reals & s < x} t; \
-\ r: Reals; #0 < r |] \
+\ r: Reals; Numeral0 < r |] \
\ ==> x + -t ~= r";
by Auto_tac;
by (forward_tac [isLubD1a RS SReal_minus] 1);
@@ -1016,7 +1016,7 @@
Goal "[| x: HFinite; \
\ isLub Reals {s. s: Reals & s < x} t; \
-\ r: Reals; #0 < r |] \
+\ r: Reals; Numeral0 < r |] \
\ ==> -(x + -t) ~= r";
by (auto_tac (claset(), simpset() addsimps [hypreal_minus_add_distrib]));
by (forward_tac [isLubD1a] 1);
@@ -1030,7 +1030,7 @@
Goal "[| x: HFinite; \
\ isLub Reals {s. s: Reals & s < x} t; \
-\ r: Reals; #0 < r |] \
+\ r: Reals; Numeral0 < r |] \
\ ==> abs (x + -t) < r";
by (forward_tac [lemma_st_part1a] 1);
by (forward_tac [lemma_st_part2a] 4);
@@ -1042,7 +1042,7 @@
Goal "[| x: HFinite; \
\ isLub Reals {s. s: Reals & s < x} t |] \
-\ ==> ALL r: Reals. #0 < r --> abs (x + -t) < r";
+\ ==> ALL r: Reals. Numeral0 < r --> abs (x + -t) < r";
by (blast_tac (claset() addSDs [lemma_st_part_major]) 1);
qed "lemma_st_part_major2";
@@ -1050,7 +1050,7 @@
Existence of real and Standard Part Theorem
----------------------------------------------*)
Goal "x: HFinite ==> \
-\ EX t: Reals. ALL r: Reals. #0 < r --> abs (x + -t) < r";
+\ EX t: Reals. ALL r: Reals. Numeral0 < r --> abs (x + -t) < r";
by (forward_tac [lemma_st_part_lub] 1 THEN Step_tac 1);
by (forward_tac [isLubD1a] 1);
by (blast_tac (claset() addDs [lemma_st_part_major2]) 1);
@@ -1089,7 +1089,7 @@
Goalw [HInfinite_def, HFinite_def] "x~: HFinite ==> x: HInfinite";
by Auto_tac;
-by (dres_inst_tac [("x","r + #1")] bspec 1);
+by (dres_inst_tac [("x","r + Numeral1")] bspec 1);
by (auto_tac (claset(), simpset() addsimps [SReal_add]));
qed "not_HFinite_HInfinite";
@@ -1241,16 +1241,16 @@
by Auto_tac;
qed "mem_monad_iff";
-Goalw [monad_def] "(x:Infinitesimal) = (x:monad #0)";
+Goalw [monad_def] "(x:Infinitesimal) = (x:monad Numeral0)";
by (auto_tac (claset() addIs [approx_sym],
simpset() addsimps [mem_infmal_iff]));
qed "Infinitesimal_monad_zero_iff";
-Goal "(x:monad #0) = (-x:monad #0)";
+Goal "(x:monad Numeral0) = (-x:monad Numeral0)";
by (simp_tac (simpset() addsimps [Infinitesimal_monad_zero_iff RS sym]) 1);
qed "monad_zero_minus_iff";
-Goal "(x:monad #0) = (abs x:monad #0)";
+Goal "(x:monad Numeral0) = (abs x:monad Numeral0)";
by (res_inst_tac [("x1","x")] (hrabs_disj RS disjE) 1);
by (auto_tac (claset(), simpset() addsimps [monad_zero_minus_iff RS sym]));
qed "monad_zero_hrabs_iff";
@@ -1286,7 +1286,7 @@
by (blast_tac (claset() addSIs [approx_sym]) 1);
qed "approx_mem_monad2";
-Goal "[| x @= y;x:monad #0 |] ==> y:monad #0";
+Goal "[| x @= y;x:monad Numeral0 |] ==> y:monad Numeral0";
by (dtac mem_monad_approx 1);
by (fast_tac (claset() addIs [approx_mem_monad,approx_trans]) 1);
qed "approx_mem_monad_zero";
@@ -1297,7 +1297,7 @@
monad_zero_hrabs_iff RS iffD1, mem_monad_approx, approx_trans3]) 1);
qed "Infinitesimal_approx_hrabs";
-Goal "[| #0 < x; x ~:Infinitesimal; e :Infinitesimal |] ==> e < x";
+Goal "[| Numeral0 < x; x ~:Infinitesimal; e :Infinitesimal |] ==> e < x";
by (rtac ccontr 1);
by (auto_tac (claset()
addIs [Infinitesimal_zero RSN (2, Infinitesimal_interval)]
@@ -1305,38 +1305,38 @@
simpset()));
qed "less_Infinitesimal_less";
-Goal "[| #0 < x; x ~: Infinitesimal; u: monad x |] ==> #0 < u";
+Goal "[| Numeral0 < x; x ~: Infinitesimal; u: monad x |] ==> Numeral0 < u";
by (dtac (mem_monad_approx RS approx_sym) 1);
by (etac (bex_Infinitesimal_iff2 RS iffD2 RS bexE) 1);
by (dres_inst_tac [("e","-xa")] less_Infinitesimal_less 1);
by Auto_tac;
qed "Ball_mem_monad_gt_zero";
-Goal "[| x < #0; x ~: Infinitesimal; u: monad x |] ==> u < #0";
+Goal "[| x < Numeral0; x ~: Infinitesimal; u: monad x |] ==> u < Numeral0";
by (dtac (mem_monad_approx RS approx_sym) 1);
by (etac (bex_Infinitesimal_iff RS iffD2 RS bexE) 1);
by (cut_inst_tac [("x","-x"),("e","xa")] less_Infinitesimal_less 1);
by Auto_tac;
qed "Ball_mem_monad_less_zero";
-Goal "[|#0 < x; x ~: Infinitesimal; x @= y|] ==> #0 < y";
+Goal "[|Numeral0 < x; x ~: Infinitesimal; x @= y|] ==> Numeral0 < y";
by (blast_tac (claset() addDs [Ball_mem_monad_gt_zero,
approx_subset_monad]) 1);
qed "lemma_approx_gt_zero";
-Goal "[|x < #0; x ~: Infinitesimal; x @= y|] ==> y < #0";
+Goal "[|x < Numeral0; x ~: Infinitesimal; x @= y|] ==> y < Numeral0";
by (blast_tac (claset() addDs [Ball_mem_monad_less_zero,
approx_subset_monad]) 1);
qed "lemma_approx_less_zero";
-Goal "[| x @= y; x < #0; x ~: Infinitesimal |] ==> abs x @= abs y";
+Goal "[| x @= y; x < Numeral0; x ~: Infinitesimal |] ==> abs x @= abs y";
by (forward_tac [lemma_approx_less_zero] 1);
by (REPEAT(assume_tac 1));
by (REPEAT(dtac hrabs_minus_eqI2 1));
by Auto_tac;
qed "approx_hrabs1";
-Goal "[| x @= y; #0 < x; x ~: Infinitesimal |] ==> abs x @= abs y";
+Goal "[| x @= y; Numeral0 < x; x ~: Infinitesimal |] ==> abs x @= abs y";
by (forward_tac [lemma_approx_gt_zero] 1);
by (REPEAT(assume_tac 1));
by (REPEAT(dtac hrabs_eqI2 1));
@@ -1345,12 +1345,12 @@
Goal "x @= y ==> abs x @= abs y";
by (res_inst_tac [("Q","x:Infinitesimal")] (excluded_middle RS disjE) 1);
-by (res_inst_tac [("x1","x"),("y1","#0")] (hypreal_linear RS disjE) 1);
+by (res_inst_tac [("x1","x"),("y1","Numeral0")] (hypreal_linear RS disjE) 1);
by (auto_tac (claset() addIs [approx_hrabs1,approx_hrabs2,
Infinitesimal_approx_hrabs], simpset()));
qed "approx_hrabs";
-Goal "abs(x) @= #0 ==> x @= #0";
+Goal "abs(x) @= Numeral0 ==> x @= Numeral0";
by (cut_inst_tac [("x","x")] hrabs_disj 1);
by (auto_tac (claset() addDs [approx_minus], simpset()));
qed "approx_hrabs_zero_cancel";
@@ -1445,7 +1445,7 @@
hypreal_of_real_le_add_Infininitesimal_cancel]) 1);
qed "hypreal_of_real_le_add_Infininitesimal_cancel2";
-Goal "[| hypreal_of_real x < e; e: Infinitesimal |] ==> hypreal_of_real x <= #0";
+Goal "[| hypreal_of_real x < e; e: Infinitesimal |] ==> hypreal_of_real x <= Numeral0";
by (rtac hypreal_leI 1 THEN Step_tac 1);
by (dtac Infinitesimal_interval 1);
by (dtac (SReal_hypreal_of_real RS SReal_Infinitesimal_zero) 4);
@@ -1453,7 +1453,7 @@
qed "hypreal_of_real_less_Infinitesimal_le_zero";
(*used once, in NSDERIV_inverse*)
-Goal "[| h: Infinitesimal; x ~= #0 |] ==> hypreal_of_real x + h ~= #0";
+Goal "[| h: Infinitesimal; x ~= Numeral0 |] ==> hypreal_of_real x + h ~= Numeral0";
by Auto_tac;
qed "Infinitesimal_add_not_zero";
@@ -1524,7 +1524,7 @@
qed "HFinite_sum_square_cancel3";
Addsimps [HFinite_sum_square_cancel3];
-Goal "[| y: monad x; #0 < hypreal_of_real e |] \
+Goal "[| y: monad x; Numeral0 < hypreal_of_real e |] \
\ ==> abs (y + -x) < hypreal_of_real e";
by (dtac (mem_monad_approx RS approx_sym) 1);
by (dtac (bex_Infinitesimal_iff RS iffD2) 1);
@@ -1682,18 +1682,18 @@
by (blast_tac (claset() addSIs [lemma_st_mult]) 1);
qed "st_mult";
-Goal "x: Infinitesimal ==> st x = #0";
+Goal "x: Infinitesimal ==> st x = Numeral0";
by (rtac (st_number_of RS subst) 1);
by (rtac approx_st_eq 1);
by (auto_tac (claset() addIs [Infinitesimal_subset_HFinite RS subsetD],
simpset() addsimps [mem_infmal_iff RS sym]));
qed "st_Infinitesimal";
-Goal "st(x) ~= #0 ==> x ~: Infinitesimal";
+Goal "st(x) ~= Numeral0 ==> x ~: Infinitesimal";
by (fast_tac (claset() addIs [st_Infinitesimal]) 1);
qed "st_not_Infinitesimal";
-Goal "[| x: HFinite; st x ~= #0 |] \
+Goal "[| x: HFinite; st x ~= Numeral0 |] \
\ ==> st(inverse x) = inverse (st x)";
by (res_inst_tac [("c1","st x")] (hypreal_mult_left_cancel RS iffD1) 1);
by (auto_tac (claset(),
@@ -1703,7 +1703,7 @@
by Auto_tac;
qed "st_inverse";
-Goal "[| x: HFinite; y: HFinite; st y ~= #0 |] \
+Goal "[| x: HFinite; y: HFinite; st y ~= Numeral0 |] \
\ ==> st(x/y) = (st x) / (st y)";
by (auto_tac (claset(),
simpset() addsimps [hypreal_divide_def, st_mult, st_not_Infinitesimal,
@@ -1747,20 +1747,20 @@
simpset() addsimps [hypreal_add_assoc RS sym]));
qed "st_le";
-Goal "[| #0 <= x; x: HFinite |] ==> #0 <= st x";
+Goal "[| Numeral0 <= x; x: HFinite |] ==> Numeral0 <= st x";
by (rtac (st_number_of RS subst) 1);
by (auto_tac (claset() addIs [st_le],
simpset() delsimps [st_number_of]));
qed "st_zero_le";
-Goal "[| x <= #0; x: HFinite |] ==> st x <= #0";
+Goal "[| x <= Numeral0; x: HFinite |] ==> st x <= Numeral0";
by (rtac (st_number_of RS subst) 1);
by (auto_tac (claset() addIs [st_le],
simpset() delsimps [st_number_of]));
qed "st_zero_ge";
Goal "x: HFinite ==> abs(st x) = st(abs x)";
-by (case_tac "#0 <= x" 1);
+by (case_tac "Numeral0 <= x" 1);
by (auto_tac (claset() addSDs [not_hypreal_leE, order_less_imp_le],
simpset() addsimps [st_zero_le,hrabs_eqI1, hrabs_minus_eqI1,
st_zero_ge,st_minus]));
@@ -1834,7 +1834,7 @@
by Auto_tac;
qed "lemma_Int_eq1";
-Goal "{n. abs (xa n) = u} <= {n. abs (xa n) < u + (#1::real)}";
+Goal "{n. abs (xa n) = u} <= {n. abs (xa n) < u + (Numeral1::real)}";
by Auto_tac;
qed "lemma_FreeUltrafilterNat_one";
@@ -1847,7 +1847,7 @@
\ |] ==> x: HFinite";
by (rtac FreeUltrafilterNat_HFinite 1);
by (res_inst_tac [("x","xa")] bexI 1);
-by (res_inst_tac [("x","u + #1")] exI 1);
+by (res_inst_tac [("x","u + Numeral1")] exI 1);
by (Ultra_tac 1 THEN assume_tac 1);
qed "FreeUltrafilterNat_const_Finite";
@@ -1915,7 +1915,7 @@
Goalw [Infinitesimal_def]
"x : Infinitesimal ==> EX X: Rep_hypreal x. \
-\ ALL u. #0 < u --> {n. abs (X n) < u}: FreeUltrafilterNat";
+\ ALL u. Numeral0 < u --> {n. abs (X n) < u}: FreeUltrafilterNat";
by (auto_tac (claset(), simpset() addsimps [hrabs_interval_iff]));
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
by (EVERY[Auto_tac, rtac bexI 1, rtac lemma_hyprel_refl 2, Step_tac 1]);
@@ -1930,7 +1930,7 @@
Goalw [Infinitesimal_def]
"EX X: Rep_hypreal x. \
-\ ALL u. #0 < u --> {n. abs (X n) < u} : FreeUltrafilterNat \
+\ ALL u. Numeral0 < u --> {n. abs (X n) < u} : FreeUltrafilterNat \
\ ==> x : Infinitesimal";
by (auto_tac (claset(),
simpset() addsimps [hrabs_interval_iff,abs_interval_iff]));
@@ -1942,7 +1942,7 @@
qed "FreeUltrafilterNat_Infinitesimal";
Goal "(x : Infinitesimal) = (EX X: Rep_hypreal x. \
-\ ALL u. #0 < u --> {n. abs (X n) < u}: FreeUltrafilterNat)";
+\ ALL u. Numeral0 < u --> {n. abs (X n) < u}: FreeUltrafilterNat)";
by (blast_tac (claset() addSIs [Infinitesimal_FreeUltrafilterNat,
FreeUltrafilterNat_Infinitesimal]) 1);
qed "Infinitesimal_FreeUltrafilterNat_iff";
@@ -1951,13 +1951,13 @@
Infinitesimals as smaller than 1/n for all n::nat (> 0)
------------------------------------------------------------------------*)
-Goal "(ALL r. #0 < r --> x < r) = (ALL n. x < inverse(real (Suc n)))";
+Goal "(ALL r. Numeral0 < r --> x < r) = (ALL n. x < inverse(real (Suc n)))";
by (auto_tac (claset(), simpset() addsimps [real_of_nat_Suc_gt_zero]));
by (blast_tac (claset() addSDs [reals_Archimedean]
addIs [order_less_trans]) 1);
qed "lemma_Infinitesimal";
-Goal "(ALL r: Reals. #0 < r --> x < r) = \
+Goal "(ALL r: Reals. Numeral0 < r --> x < r) = \
\ (ALL n. x < inverse(hypreal_of_nat (Suc n)))";
by (Step_tac 1);
by (dres_inst_tac [("x","inverse (hypreal_of_real(real (Suc n)))")]
@@ -2089,7 +2089,7 @@
qed "HFinite_epsilon";
Addsimps [HFinite_epsilon];
-Goal "epsilon @= #0";
+Goal "epsilon @= Numeral0";
by (simp_tac (simpset() addsimps [mem_infmal_iff RS sym]) 1);
qed "epsilon_approx_zero";
Addsimps [epsilon_approx_zero];
@@ -2109,7 +2109,7 @@
by (simp_tac (simpset() addsimps [real_of_nat_Suc_gt_zero]) 1);
qed "real_of_nat_less_inverse_iff";
-Goal "#0 < u ==> finite {n. u < inverse(real(Suc n))}";
+Goal "Numeral0 < u ==> finite {n. u < inverse(real(Suc n))}";
by (asm_simp_tac (simpset() addsimps [real_of_nat_less_inverse_iff]) 1);
by (asm_simp_tac (simpset() addsimps [real_of_nat_Suc,
real_less_diff_eq RS sym]) 1);
@@ -2122,7 +2122,7 @@
simpset() addsimps [order_less_imp_le]));
qed "lemma_real_le_Un_eq2";
-Goal "(inverse (real(Suc n)) <= r) = (#1 <= r * real(Suc n))";
+Goal "(inverse (real(Suc n)) <= r) = (Numeral1 <= r * real(Suc n))";
by (simp_tac (simpset() addsimps [linorder_not_less RS sym]) 1);
by (simp_tac (simpset() addsimps [real_inverse_eq_divide]) 1);
by (stac pos_real_less_divide_eq 1);
@@ -2138,18 +2138,18 @@
Goal "finite {n::nat. u = inverse(real(Suc n))}";
by (asm_simp_tac (simpset() addsimps [real_of_nat_inverse_eq_iff]) 1);
-by (cut_inst_tac [("x","inverse u - #1")] lemma_finite_omega_set 1);
+by (cut_inst_tac [("x","inverse u - Numeral1")] lemma_finite_omega_set 1);
by (asm_full_simp_tac (simpset() addsimps [real_of_nat_Suc,
real_diff_eq_eq RS sym, eq_commute]) 1);
qed "lemma_finite_omega_set2";
-Goal "#0 < u ==> finite {n. u <= inverse(real(Suc n))}";
+Goal "Numeral0 < u ==> finite {n. u <= inverse(real(Suc n))}";
by (auto_tac (claset(),
simpset() addsimps [lemma_real_le_Un_eq2,lemma_finite_omega_set2,
finite_inverse_real_of_posnat_gt_real]));
qed "finite_inverse_real_of_posnat_ge_real";
-Goal "#0 < u ==> \
+Goal "Numeral0 < u ==> \
\ {n. u <= inverse(real(Suc n))} ~: FreeUltrafilterNat";
by (blast_tac (claset() addSIs [FreeUltrafilterNat_finite,
finite_inverse_real_of_posnat_ge_real]) 1);
@@ -2166,7 +2166,7 @@
simpset() addsimps [not_real_leE]));
val lemma = result();
-Goal "#0 < u ==> \
+Goal "Numeral0 < u ==> \
\ {n. inverse(real(Suc n)) < u} : FreeUltrafilterNat";
by (cut_inst_tac [("u","u")] inverse_real_of_posnat_ge_real_FreeUltrafilterNat 1);
by (auto_tac (claset() addDs [FreeUltrafilterNat_Compl_mem],
--- a/src/HOL/Hyperreal/NatStar.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Hyperreal/NatStar.ML Fri Oct 05 21:52:39 2001 +0200
@@ -404,7 +404,7 @@
Goal "N : HNatInfinite \
\ ==> (*fNat* (%x::nat. inverse(real x))) N = inverse(hypreal_of_hypnat N)";
by (res_inst_tac [("f1","inverse")] (starfun_stafunNat_o2 RS subst) 1);
-by (subgoal_tac "hypreal_of_hypnat N ~= #0" 1);
+by (subgoal_tac "hypreal_of_hypnat N ~= Numeral0" 1);
by (auto_tac (claset(),
simpset() addsimps [starfunNat_real_of_nat, starfun_inverse_inverse]));
qed "starfunNat_inverse_real_of_nat_eq";
--- a/src/HOL/Hyperreal/SEQ.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Hyperreal/SEQ.ML Fri Oct 05 21:52:39 2001 +0200
@@ -26,7 +26,7 @@
Goalw [LIMSEQ_def]
"(X ----> L) = \
-\ (ALL r. #0 <r --> (EX no. ALL n. no <= n --> abs(X n + -L) < r))";
+\ (ALL r. Numeral0 <r --> (EX no. ALL n. no <= n --> abs(X n + -L) < r))";
by (Simp_tac 1);
qed "LIMSEQ_iff";
@@ -120,7 +120,7 @@
by Auto_tac;
val lemmaLIM2 = result();
-Goal "[| #0 < r; ALL n. r <= abs (X (f n) + - L); \
+Goal "[| Numeral0 < r; ALL n. r <= abs (X (f n) + - L); \
\ (*fNat* X) (Abs_hypnat (hypnatrel `` {f})) + \
\ - hypreal_of_real L @= 0 |] ==> False";
by (auto_tac (claset(),simpset() addsimps [starfunNat,
@@ -234,7 +234,7 @@
Proof is like that of NSLIM_inverse.
--------------------------------------------------------------*)
Goalw [NSLIMSEQ_def]
- "[| X ----NS> a; a ~= #0 |] ==> (%n. inverse(X n)) ----NS> inverse(a)";
+ "[| X ----NS> a; a ~= Numeral0 |] ==> (%n. inverse(X n)) ----NS> inverse(a)";
by (Clarify_tac 1);
by (dtac bspec 1);
by (auto_tac (claset(),
@@ -244,18 +244,18 @@
(*------ Standard version of theorem -------*)
-Goal "[| X ----> a; a ~= #0 |] ==> (%n. inverse(X n)) ----> inverse(a)";
+Goal "[| X ----> a; a ~= Numeral0 |] ==> (%n. inverse(X n)) ----> inverse(a)";
by (asm_full_simp_tac (simpset() addsimps [NSLIMSEQ_inverse,
LIMSEQ_NSLIMSEQ_iff]) 1);
qed "LIMSEQ_inverse";
-Goal "[| X ----NS> a; Y ----NS> b; b ~= #0 |] \
+Goal "[| X ----NS> a; Y ----NS> b; b ~= Numeral0 |] \
\ ==> (%n. X n / Y n) ----NS> a/b";
by (asm_full_simp_tac (simpset() addsimps [NSLIMSEQ_mult, NSLIMSEQ_inverse,
real_divide_def]) 1);
qed "NSLIMSEQ_mult_inverse";
-Goal "[| X ----> a; Y ----> b; b ~= #0 |] ==> (%n. X n / Y n) ----> a/b";
+Goal "[| X ----> a; Y ----> b; b ~= Numeral0 |] ==> (%n. X n / Y n) ----> a/b";
by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_mult, LIMSEQ_inverse,
real_divide_def]) 1);
qed "LIMSEQ_divide";
@@ -376,16 +376,16 @@
Bounded Sequence
------------------------------------------------------------------*)
Goalw [Bseq_def]
- "Bseq X ==> EX K. #0 < K & (ALL n. abs(X n) <= K)";
+ "Bseq X ==> EX K. Numeral0 < K & (ALL n. abs(X n) <= K)";
by (assume_tac 1);
qed "BseqD";
Goalw [Bseq_def]
- "[| #0 < K; ALL n. abs(X n) <= K |] ==> Bseq X";
+ "[| Numeral0 < K; ALL n. abs(X n) <= K |] ==> Bseq X";
by (Blast_tac 1);
qed "BseqI";
-Goal "(EX K. #0 < K & (ALL n. abs(X n) <= K)) = \
+Goal "(EX K. Numeral0 < K & (ALL n. abs(X n) <= K)) = \
\ (EX N. ALL n. abs(X n) <= real(Suc N))";
by Auto_tac;
by (cut_inst_tac [("x","K")] reals_Archimedean2 1);
@@ -401,7 +401,7 @@
by (simp_tac (simpset() addsimps [lemma_NBseq_def]) 1);
qed "Bseq_iff";
-Goal "(EX K. #0 < K & (ALL n. abs(X n) <= K)) = \
+Goal "(EX K. Numeral0 < K & (ALL n. abs(X n) <= K)) = \
\ (EX N. ALL n. abs(X n) < real(Suc N))";
by (stac lemma_NBseq_def 1);
by Auto_tac;
@@ -444,7 +444,7 @@
HNatInfinite_FreeUltrafilterNat_iff]));
by (EVERY[rtac bexI 1, rtac lemma_hyprel_refl 2]);
by (dres_inst_tac [("f","Xa")] lemma_Bseq 1);
-by (res_inst_tac [("x","K+#1")] exI 1);
+by (res_inst_tac [("x","K+Numeral1")] exI 1);
by (rotate_tac 2 1 THEN dtac FreeUltrafilterNat_all 1);
by (Ultra_tac 1);
qed "Bseq_NSBseq";
@@ -461,14 +461,14 @@
is not what we want (read useless!)
-------------------------------------------------------------------*)
-Goal "ALL K. #0 < K --> (EX n. K < abs (X n)) \
+Goal "ALL K. Numeral0 < K --> (EX n. K < abs (X n)) \
\ ==> ALL N. EX n. real(Suc N) < abs (X n)";
by (Step_tac 1);
by (cut_inst_tac [("n","N")] real_of_nat_Suc_gt_zero 1);
by (Blast_tac 1);
val lemmaNSBseq = result();
-Goal "ALL K. #0 < K --> (EX n. K < abs (X n)) \
+Goal "ALL K. Numeral0 < K --> (EX n. K < abs (X n)) \
\ ==> EX f. ALL N. real(Suc N) < abs (X (f N))";
by (dtac lemmaNSBseq 1);
by (dtac choice 1);
@@ -652,7 +652,7 @@
Goal "!!(X::nat=> real). \
\ [| ALL m. X m ~= U; \
\ isLub UNIV {x. EX n. X n = x} U; \
-\ #0 < T; \
+\ Numeral0 < T; \
\ U + - T < U \
\ |] ==> EX m. U + -T < X m & X m < U";
by (dtac lemma_converg2 1 THEN assume_tac 1);
@@ -722,7 +722,7 @@
(***--- alternative formulation for boundedness---***)
Goalw [Bseq_def]
- "Bseq X = (EX k x. #0 < k & (ALL n. abs(X(n) + -x) <= k))";
+ "Bseq X = (EX k x. Numeral0 < k & (ALL n. abs(X(n) + -x) <= k))";
by (Step_tac 1);
by (res_inst_tac [("x","k + abs(x)")] exI 2);
by (res_inst_tac [("x","K")] exI 1);
@@ -733,7 +733,7 @@
qed "Bseq_iff2";
(***--- alternative formulation for boundedness ---***)
-Goal "Bseq X = (EX k N. #0 < k & (ALL n. abs(X(n) + -X(N)) <= k))";
+Goal "Bseq X = (EX k N. Numeral0 < k & (ALL n. abs(X(n) + -X(N)) <= k))";
by (Step_tac 1);
by (asm_full_simp_tac (simpset() addsimps [Bseq_def]) 1);
by (Step_tac 1);
@@ -748,7 +748,7 @@
qed "Bseq_iff3";
Goalw [Bseq_def] "(ALL n. k <= f n & f n <= K) ==> Bseq f";
-by (res_inst_tac [("x","(abs(k) + abs(K)) + #1")] exI 1);
+by (res_inst_tac [("x","(abs(k) + abs(K)) + Numeral1")] exI 1);
by (Auto_tac);
by (dres_inst_tac [("x","n")] spec 2);
by (ALLGOALS arith_tac);
@@ -841,8 +841,8 @@
-------------------------------------------------------*)
(***------------- VARIOUS LEMMAS --------------***)
-Goal "ALL n. M <= n --> abs (X M + - X n) < (#1::real) \
-\ ==> ALL n. M <= n --> abs(X n) < #1 + abs(X M)";
+Goal "ALL n. M <= n --> abs (X M + - X n) < (Numeral1::real) \
+\ ==> ALL n. M <= n --> abs(X n) < Numeral1 + abs(X M)";
by (Step_tac 1);
by (dtac spec 1 THEN Auto_tac);
by (arith_tac 1);
@@ -911,7 +911,7 @@
outlines sketched by various authors would suggest
---------------------------------------------------------*)
Goalw [Cauchy_def,Bseq_def] "Cauchy X ==> Bseq X";
-by (dres_inst_tac [("x","#1")] spec 1);
+by (dres_inst_tac [("x","Numeral1")] spec 1);
by (etac (rename_numerals real_zero_less_one RSN (2,impE)) 1);
by (Step_tac 1);
by (dres_inst_tac [("x","M")] spec 1);
@@ -920,7 +920,7 @@
by (cut_inst_tac [("M","M"),("X","X")] SUP_rabs_subseq 1);
by (Step_tac 1);
by (cut_inst_tac [("R1.0","abs(X m)"),
- ("R2.0","#1 + abs(X M)")] real_linear 1);
+ ("R2.0","Numeral1 + abs(X M)")] real_linear 1);
by (Step_tac 1);
by (dtac lemma_trans1 1 THEN assume_tac 1);
by (dtac lemma_trans2 3 THEN assume_tac 3);
@@ -928,8 +928,8 @@
by (dtac (abs_add_one_gt_zero RS order_less_trans) 3);
by (dtac lemma_trans4 1);
by (dtac lemma_trans4 2);
-by (res_inst_tac [("x","#1 + abs(X M)")] exI 1);
-by (res_inst_tac [("x","#1 + abs(X M)")] exI 2);
+by (res_inst_tac [("x","Numeral1 + abs(X M)")] exI 1);
+by (res_inst_tac [("x","Numeral1 + abs(X M)")] exI 2);
by (res_inst_tac [("x","abs(X m)")] exI 3);
by (auto_tac (claset() addSEs [lemma_Nat_covered],
simpset()));
@@ -1082,7 +1082,7 @@
----------------------------------------------------*)
(* we can prove this directly since proof is trivial *)
Goalw [LIMSEQ_def]
- "((%n. abs(f n)) ----> #0) = (f ----> #0)";
+ "((%n. abs(f n)) ----> Numeral0) = (f ----> Numeral0)";
by (simp_tac (simpset() addsimps [abs_idempotent]) 1);
qed "LIMSEQ_rabs_zero";
@@ -1092,7 +1092,7 @@
(* than the direct standard one above! *)
(*-----------------------------------------------------*)
-Goal "((%n. abs(f n)) ----NS> #0) = (f ----NS> #0)";
+Goal "((%n. abs(f n)) ----NS> Numeral0) = (f ----NS> Numeral0)";
by (simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff RS sym,
LIMSEQ_rabs_zero]) 1);
qed "NSLIMSEQ_rabs_zero";
@@ -1119,7 +1119,7 @@
(* standard proof seems easier *)
Goalw [LIMSEQ_def]
"ALL y. EX N. ALL n. N <= n --> y < f(n) \
-\ ==> (%n. inverse(f n)) ----> #0";
+\ ==> (%n. inverse(f n)) ----> Numeral0";
by (Step_tac 1 THEN Asm_full_simp_tac 1);
by (dres_inst_tac [("x","inverse r")] spec 1 THEN Step_tac 1);
by (res_inst_tac [("x","N")] exI 1 THEN Step_tac 1);
@@ -1134,7 +1134,7 @@
qed "LIMSEQ_inverse_zero";
Goal "ALL y. EX N. ALL n. N <= n --> y < f(n) \
-\ ==> (%n. inverse(f n)) ----NS> #0";
+\ ==> (%n. inverse(f n)) ----NS> Numeral0";
by (asm_simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff RS sym,
LIMSEQ_inverse_zero]) 1);
qed "NSLIMSEQ_inverse_zero";
@@ -1143,7 +1143,7 @@
Sequence 1/n --> 0 as n --> infinity
-------------------------------------------------------------*)
-Goal "(%n. inverse(real(Suc n))) ----> #0";
+Goal "(%n. inverse(real(Suc n))) ----> Numeral0";
by (rtac LIMSEQ_inverse_zero 1 THEN Step_tac 1);
by (cut_inst_tac [("x","y")] reals_Archimedean2 1);
by (Step_tac 1 THEN res_inst_tac [("x","n")] exI 1);
@@ -1153,7 +1153,7 @@
by (blast_tac (claset() addIs [order_less_le_trans]) 1);
qed "LIMSEQ_inverse_real_of_nat";
-Goal "(%n. inverse(real(Suc n))) ----NS> #0";
+Goal "(%n. inverse(real(Suc n))) ----NS> Numeral0";
by (simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff RS sym,
LIMSEQ_inverse_real_of_nat]) 1);
qed "NSLIMSEQ_inverse_real_of_nat";
@@ -1188,13 +1188,13 @@
LIMSEQ_inverse_real_of_posnat_add_minus]) 1);
qed "NSLIMSEQ_inverse_real_of_posnat_add_minus";
-Goal "(%n. r*( #1 + -inverse(real(Suc n)))) ----> r";
-by (cut_inst_tac [("b","#1")] ([LIMSEQ_const,
+Goal "(%n. r*( Numeral1 + -inverse(real(Suc n)))) ----> r";
+by (cut_inst_tac [("b","Numeral1")] ([LIMSEQ_const,
LIMSEQ_inverse_real_of_posnat_add_minus] MRS LIMSEQ_mult) 1);
by (Auto_tac);
qed "LIMSEQ_inverse_real_of_posnat_add_minus_mult";
-Goal "(%n. r*( #1 + -inverse(real(Suc n)))) ----NS> r";
+Goal "(%n. r*( Numeral1 + -inverse(real(Suc n)))) ----NS> r";
by (simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff RS sym,
LIMSEQ_inverse_real_of_posnat_add_minus_mult]) 1);
qed "NSLIMSEQ_inverse_real_of_posnat_add_minus_mult";
@@ -1214,22 +1214,22 @@
qed "LIMSEQ_pow";
(*----------------------------------------------------------------
- 0 <= x < #1 ==> (x ^ n ----> 0)
+ 0 <= x < Numeral1 ==> (x ^ n ----> 0)
Proof will use (NS) Cauchy equivalence for convergence and
also fact that bounded and monotonic sequence converges.
---------------------------------------------------------------*)
-Goalw [Bseq_def] "[| #0 <= x; x < #1 |] ==> Bseq (%n. x ^ n)";
-by (res_inst_tac [("x","#1")] exI 1);
+Goalw [Bseq_def] "[| Numeral0 <= x; x < Numeral1 |] ==> Bseq (%n. x ^ n)";
+by (res_inst_tac [("x","Numeral1")] exI 1);
by (auto_tac (claset() addDs [conjI RS realpow_le]
addIs [order_less_imp_le],
simpset() addsimps [abs_eqI1, realpow_abs RS sym] ));
qed "Bseq_realpow";
-Goal "[| #0 <= x; x < #1 |] ==> monoseq (%n. x ^ n)";
+Goal "[| Numeral0 <= x; x < Numeral1 |] ==> monoseq (%n. x ^ n)";
by (blast_tac (claset() addSIs [mono_SucI2,realpow_Suc_le3]) 1);
qed "monoseq_realpow";
-Goal "[| #0 <= x; x < #1 |] ==> convergent (%n. x ^ n)";
+Goal "[| Numeral0 <= x; x < Numeral1 |] ==> convergent (%n. x ^ n)";
by (blast_tac (claset() addSIs [Bseq_monoseq_convergent,
Bseq_realpow,monoseq_realpow]) 1);
qed "convergent_realpow";
@@ -1238,7 +1238,7 @@
Goalw [NSLIMSEQ_def]
- "[| #0 <= x; x < #1 |] ==> (%n. x ^ n) ----NS> #0";
+ "[| Numeral0 <= x; x < Numeral1 |] ==> (%n. x ^ n) ----NS> Numeral0";
by (auto_tac (claset() addSDs [convergent_realpow],
simpset() addsimps [convergent_NSconvergent_iff]));
by (forward_tac [NSconvergentD] 1);
@@ -1258,12 +1258,12 @@
qed "NSLIMSEQ_realpow_zero";
(*--------------- standard version ---------------*)
-Goal "[| #0 <= x; x < #1 |] ==> (%n. x ^ n) ----> #0";
+Goal "[| Numeral0 <= x; x < Numeral1 |] ==> (%n. x ^ n) ----> Numeral0";
by (asm_simp_tac (simpset() addsimps [NSLIMSEQ_realpow_zero,
LIMSEQ_NSLIMSEQ_iff]) 1);
qed "LIMSEQ_realpow_zero";
-Goal "#1 < x ==> (%n. a / (x ^ n)) ----> #0";
+Goal "Numeral1 < x ==> (%n. a / (x ^ n)) ----> Numeral0";
by (cut_inst_tac [("a","a"),("x1","inverse x")]
([LIMSEQ_const, LIMSEQ_realpow_zero] MRS LIMSEQ_mult) 1);
by (auto_tac (claset(),
@@ -1275,22 +1275,22 @@
(*----------------------------------------------------------------
Limit of c^n for |c| < 1
---------------------------------------------------------------*)
-Goal "abs(c) < #1 ==> (%n. abs(c) ^ n) ----> #0";
+Goal "abs(c) < Numeral1 ==> (%n. abs(c) ^ n) ----> Numeral0";
by (blast_tac (claset() addSIs [LIMSEQ_realpow_zero,abs_ge_zero]) 1);
qed "LIMSEQ_rabs_realpow_zero";
-Goal "abs(c) < #1 ==> (%n. abs(c) ^ n) ----NS> #0";
+Goal "abs(c) < Numeral1 ==> (%n. abs(c) ^ n) ----NS> Numeral0";
by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_rabs_realpow_zero,
LIMSEQ_NSLIMSEQ_iff RS sym]) 1);
qed "NSLIMSEQ_rabs_realpow_zero";
-Goal "abs(c) < #1 ==> (%n. c ^ n) ----> #0";
+Goal "abs(c) < Numeral1 ==> (%n. c ^ n) ----> Numeral0";
by (rtac (LIMSEQ_rabs_zero RS iffD1) 1);
by (auto_tac (claset() addIs [LIMSEQ_rabs_realpow_zero],
simpset() addsimps [realpow_abs RS sym]));
qed "LIMSEQ_rabs_realpow_zero2";
-Goal "abs(c) < #1 ==> (%n. c ^ n) ----NS> #0";
+Goal "abs(c) < Numeral1 ==> (%n. c ^ n) ----NS> Numeral0";
by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_rabs_realpow_zero2,
LIMSEQ_NSLIMSEQ_iff RS sym]) 1);
qed "NSLIMSEQ_rabs_realpow_zero2";
@@ -1308,7 +1308,7 @@
(*** A sequence converging to zero defines an infinitesimal ***)
Goalw [NSLIMSEQ_def]
- "X ----NS> #0 ==> Abs_hypreal(hyprel``{X}) : Infinitesimal";
+ "X ----NS> Numeral0 ==> Abs_hypreal(hyprel``{X}) : Infinitesimal";
by (dres_inst_tac [("x","whn")] bspec 1);
by (simp_tac (simpset() addsimps [HNatInfinite_whn]) 1);
by (auto_tac (claset(),
--- a/src/HOL/Hyperreal/SEQ.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Hyperreal/SEQ.thy Fri Oct 05 21:52:39 2001 +0200
@@ -10,7 +10,7 @@
(* Standard definition of convergence of sequence *)
LIMSEQ :: [nat=>real,real] => bool ("((_)/ ----> (_))" [60, 60] 60)
- "X ----> L == (ALL r. #0 < r --> (EX no. ALL n. no <= n --> abs (X n + -L) < r))"
+ "X ----> L == (ALL r. Numeral0 < r --> (EX no. ALL n. no <= n --> abs (X n + -L) < r))"
(* Nonstandard definition of convergence of sequence *)
NSLIMSEQ :: [nat=>real,real] => bool ("((_)/ ----NS> (_))" [60, 60] 60)
@@ -33,7 +33,7 @@
(* Standard definition for bounded sequence *)
Bseq :: (nat => real) => bool
- "Bseq X == (EX K. (#0 < K & (ALL n. abs(X n) <= K)))"
+ "Bseq X == (EX K. (Numeral0 < K & (ALL n. abs(X n) <= K)))"
(* Nonstandard definition for bounded sequence *)
NSBseq :: (nat=>real) => bool
@@ -52,7 +52,7 @@
(* Standard definition *)
Cauchy :: (nat => real) => bool
- "Cauchy X == (ALL e. (#0 < e -->
+ "Cauchy X == (ALL e. (Numeral0 < e -->
(EX M. (ALL m n. M <= m & M <= n
--> abs((X m) + -(X n)) < e))))"
--- a/src/HOL/Hyperreal/Series.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Hyperreal/Series.ML Fri Oct 05 21:52:39 2001 +0200
@@ -5,13 +5,13 @@
Description : Finite summation and infinite series
*)
-Goal "sumr (Suc n) n f = #0";
+Goal "sumr (Suc n) n f = Numeral0";
by (induct_tac "n" 1);
by (Auto_tac);
qed "sumr_Suc_zero";
Addsimps [sumr_Suc_zero];
-Goal "sumr m m f = #0";
+Goal "sumr m m f = Numeral0";
by (induct_tac "m" 1);
by (Auto_tac);
qed "sumr_eq_bounds";
@@ -22,7 +22,7 @@
qed "sumr_Suc_eq";
Addsimps [sumr_Suc_eq];
-Goal "sumr (m+k) k f = #0";
+Goal "sumr (m+k) k f = Numeral0";
by (induct_tac "k" 1);
by (Auto_tac);
qed "sumr_add_lbound_zero";
@@ -83,7 +83,7 @@
by (full_simp_tac (simpset() addsimps [sumr_add_mult_const]) 1);
qed "sumr_diff_mult_const";
-Goal "n < m --> sumr m n f = #0";
+Goal "n < m --> sumr m n f = Numeral0";
by (induct_tac "n" 1);
by (auto_tac (claset() addDs [less_imp_le], simpset()));
qed_spec_mp "sumr_less_bounds_zero";
@@ -101,7 +101,7 @@
by (Auto_tac);
qed "sumr_shift_bounds";
-Goal "sumr 0 (#2*n) (%i. (#-1) ^ Suc i) = #0";
+Goal "sumr 0 (# 2*n) (%i. (# -1) ^ Suc i) = Numeral0";
by (induct_tac "n" 1);
by (Auto_tac);
qed "sumr_minus_one_realpow_zero";
@@ -137,7 +137,7 @@
real_of_nat_Suc]) 1);
qed_spec_mp "sumr_interval_const2";
-Goal "(ALL n. m <= n --> #0 <= f n) & m < k --> sumr 0 m f <= sumr 0 k f";
+Goal "(ALL n. m <= n --> Numeral0 <= f n) & m < k --> sumr 0 m f <= sumr 0 k f";
by (induct_tac "k" 1);
by (Step_tac 1);
by (ALLGOALS(asm_full_simp_tac (simpset() addsimps [less_Suc_eq_le])));
@@ -156,21 +156,21 @@
simpset() addsimps [le_def]));
qed_spec_mp "sumr_le2";
-Goal "(ALL n. #0 <= f n) --> #0 <= sumr m n f";
+Goal "(ALL n. Numeral0 <= f n) --> Numeral0 <= sumr m n f";
by (induct_tac "n" 1);
by Auto_tac;
by (dres_inst_tac [("x","n")] spec 1);
by (arith_tac 1);
qed_spec_mp "sumr_ge_zero";
-Goal "(ALL n. m <= n --> #0 <= f n) --> #0 <= sumr m n f";
+Goal "(ALL n. m <= n --> Numeral0 <= f n) --> Numeral0 <= sumr m n f";
by (induct_tac "n" 1);
by Auto_tac;
by (dres_inst_tac [("x","n")] spec 1);
by (arith_tac 1);
qed_spec_mp "sumr_ge_zero2";
-Goal "#0 <= sumr m n (%n. abs (f n))";
+Goal "Numeral0 <= sumr m n (%n. abs (f n))";
by (induct_tac "n" 1);
by Auto_tac;
by (arith_tac 1);
@@ -184,21 +184,21 @@
qed "rabs_sumr_rabs_cancel";
Addsimps [rabs_sumr_rabs_cancel];
-Goal "ALL n. N <= n --> f n = #0 \
-\ ==> ALL m n. N <= m --> sumr m n f = #0";
+Goal "ALL n. N <= n --> f n = Numeral0 \
+\ ==> ALL m n. N <= m --> sumr m n f = Numeral0";
by (Step_tac 1);
by (induct_tac "n" 1);
by (Auto_tac);
qed "sumr_zero";
-Goal "ALL n. N <= n --> f (Suc n) = #0 \
-\ ==> ALL m n. Suc N <= m --> sumr m n f = #0";
+Goal "ALL n. N <= n --> f (Suc n) = Numeral0 \
+\ ==> ALL m n. Suc N <= m --> sumr m n f = Numeral0";
by (rtac sumr_zero 1 THEN Step_tac 1);
by (case_tac "n" 1);
by Auto_tac;
qed "Suc_le_imp_diff_ge2";
-Goal "sumr 1' n (%n. f(n) * #0 ^ n) = #0";
+Goal "sumr (Suc 0) n (%n. f(n) * Numeral0 ^ n) = Numeral0";
by (induct_tac "n" 1);
by (case_tac "n" 2);
by Auto_tac;
@@ -269,7 +269,7 @@
(*
Goalw [sums_def,LIMSEQ_def]
- "(ALL m. n <= Suc m --> f(m) = #0) ==> f sums (sumr 0 n f)";
+ "(ALL m. n <= Suc m --> f(m) = Numeral0) ==> f sums (sumr 0 n f)";
by (Step_tac 1);
by (res_inst_tac [("x","n")] exI 1);
by (Step_tac 1 THEN forward_tac [le_imp_less_or_eq] 1);
@@ -283,7 +283,7 @@
**********************)
Goalw [sums_def,LIMSEQ_def]
- "(ALL m. n <= m --> f(m) = #0) ==> f sums (sumr 0 n f)";
+ "(ALL m. n <= m --> f(m) = Numeral0) ==> f sums (sumr 0 n f)";
by (Step_tac 1);
by (res_inst_tac [("x","n")] exI 1);
by (Step_tac 1 THEN forward_tac [le_imp_less_or_eq] 1);
@@ -341,35 +341,35 @@
by (Auto_tac);
qed "sums_group";
-Goal "[|summable f; ALL d. #0 < (f(n + (2 * d))) + f(n + ((2 * d) + 1))|] \
+Goal "[|summable f; ALL d. Numeral0 < (f(n + (Suc (Suc 0) * d))) + f(n + ((Suc (Suc 0) * d) + 1))|] \
\ ==> sumr 0 n f < suminf f";
by (dtac summable_sums 1);
by (auto_tac (claset(),simpset() addsimps [sums_def,LIMSEQ_def]));
by (dres_inst_tac [("x","f(n) + f(n + 1)")] spec 1);
by (Auto_tac);
by (rtac ccontr 2 THEN dtac real_leI 2);
-by (subgoal_tac "sumr 0 (n + 2) f <= sumr 0 (2 * (Suc no) + n) f" 2);
+by (subgoal_tac "sumr 0 (n + Suc (Suc 0)) f <= sumr 0 (Suc (Suc 0) * (Suc no) + n) f" 2);
by (induct_tac "no" 3 THEN Simp_tac 3);
-by (res_inst_tac [("y","sumr 0 (2*(Suc na)+n) f")] order_trans 3);
+by (res_inst_tac [("y","sumr 0 (Suc (Suc 0)*(Suc na)+n) f")] order_trans 3);
by (assume_tac 3);
by (dres_inst_tac [("x","Suc na")] spec 3);
by (dres_inst_tac [("x","0")] spec 1);
by (Asm_full_simp_tac 1);
by (asm_full_simp_tac (simpset() addsimps add_ac) 2);
-by (rotate_tac 1 1 THEN dres_inst_tac [("x","2 * (Suc no) + n")] spec 1);
+by (rotate_tac 1 1 THEN dres_inst_tac [("x","Suc (Suc 0) * (Suc no) + n")] spec 1);
by (Step_tac 1 THEN Asm_full_simp_tac 1);
by (subgoal_tac "suminf f + (f(n) + f(n + 1)) <= \
-\ sumr 0 (2 * (Suc no) + n) f" 1);
-by (res_inst_tac [("y","sumr 0 (n+2) f")] order_trans 2);
+\ sumr 0 (Suc (Suc 0) * (Suc no) + n) f" 1);
+by (res_inst_tac [("y","sumr 0 (n+ Suc (Suc 0)) f")] order_trans 2);
by (assume_tac 3);
by (res_inst_tac [("y","sumr 0 n f + (f(n) + f(n + 1))")] order_trans 2);
by (REPEAT(Asm_simp_tac 2));
-by (subgoal_tac "suminf f <= sumr 0 (2 * (Suc no) + n) f" 1);
+by (subgoal_tac "suminf f <= sumr 0 (Suc (Suc 0) * (Suc no) + n) f" 1);
by (res_inst_tac [("y","suminf f + (f(n) + f(n + 1))")] order_trans 2);
by (assume_tac 3);
by (dres_inst_tac [("x","0")] spec 2);
by (Asm_full_simp_tac 2);
-by (subgoal_tac "#0 <= sumr 0 (2 * Suc no + n) f + - suminf f" 1);
+by (subgoal_tac "Numeral0 <= sumr 0 (Suc (Suc 0) * Suc no + n) f + - suminf f" 1);
by (dtac (rename_numerals abs_eqI1) 1 );
by (Asm_full_simp_tac 1);
by (auto_tac (claset(),simpset() addsimps [real_le_def]));
@@ -379,7 +379,7 @@
Summable series of positive terms has limit >= any partial sum
----------------------------------------------------------------*)
Goal
- "[| summable f; ALL m. n <= m --> #0 <= f(m) |] \
+ "[| summable f; ALL m. n <= m --> Numeral0 <= f(m) |] \
\ ==> sumr 0 n f <= suminf f";
by (dtac summable_sums 1);
by (rewtac sums_def);
@@ -390,7 +390,7 @@
by (auto_tac (claset() addIs [sumr_le], simpset()));
qed "series_pos_le";
-Goal "[| summable f; ALL m. n <= m --> #0 < f(m) |] \
+Goal "[| summable f; ALL m. n <= m --> Numeral0 < f(m) |] \
\ ==> sumr 0 n f < suminf f";
by (res_inst_tac [("y","sumr 0 (Suc n) f")] order_less_le_trans 1);
by (rtac series_pos_le 2);
@@ -403,10 +403,10 @@
sum of geometric progression
-------------------------------------------------------------------*)
-Goal "x ~= #1 ==> sumr 0 n (%n. x ^ n) = (x ^ n - #1) / (x - #1)";
+Goal "x ~= Numeral1 ==> sumr 0 n (%n. x ^ n) = (x ^ n - Numeral1) / (x - Numeral1)";
by (induct_tac "n" 1);
by (Auto_tac);
-by (res_inst_tac [("c1","x - #1")] (real_mult_right_cancel RS iffD1) 1);
+by (res_inst_tac [("c1","x - Numeral1")] (real_mult_right_cancel RS iffD1) 1);
by (auto_tac (claset(),
simpset() addsimps [real_mult_assoc, real_add_mult_distrib]));
by (auto_tac (claset(),
@@ -414,12 +414,12 @@
real_diff_def, real_mult_commute]));
qed "sumr_geometric";
-Goal "abs(x) < #1 ==> (%n. x ^ n) sums (#1/(#1 - x))";
-by (case_tac "x = #1" 1);
+Goal "abs(x) < Numeral1 ==> (%n. x ^ n) sums (Numeral1/(Numeral1 - x))";
+by (case_tac "x = Numeral1" 1);
by (auto_tac (claset() addSDs [LIMSEQ_rabs_realpow_zero2],
simpset() addsimps [sumr_geometric ,sums_def,
real_diff_def, real_add_divide_distrib]));
-by (subgoal_tac "#1 / (#1 + - x) = #0/(x-#1) + - #1/(x-#1)" 1);
+by (subgoal_tac "Numeral1 / (Numeral1 + - x) = Numeral0/(x-Numeral1) + - Numeral1/(x-Numeral1)" 1);
by (asm_full_simp_tac (simpset() addsimps [real_divide_eq_cancel1,
real_divide_minus_eq RS sym, real_diff_def]) 2);
by (etac ssubst 1);
@@ -437,7 +437,7 @@
qed "summable_convergent_sumr_iff";
Goal "summable f = \
-\ (ALL e. #0 < e --> (EX N. ALL m n. N <= m --> abs(sumr m n f) < e))";
+\ (ALL e. Numeral0 < e --> (EX N. ALL m n. N <= m --> abs(sumr m n f) < e))";
by (auto_tac (claset(),simpset() addsimps [summable_convergent_sumr_iff,
Cauchy_convergent_iff RS sym,Cauchy_def]));
by (ALLGOALS(dtac spec) THEN Auto_tac);
@@ -455,7 +455,7 @@
(*-------------------------------------------------------------------
Terms of a convergent series tend to zero
- > Goalw [LIMSEQ_def] "summable f ==> f ----> #0";
+ > Goalw [LIMSEQ_def] "summable f ==> f ----> Numeral0";
Proved easily in HSeries after proving nonstandard case.
-------------------------------------------------------------------*)
(*-------------------------------------------------------------------
@@ -527,10 +527,10 @@
The ratio test
-------------------------------------------------------------------*)
-Goal "[| c <= #0; abs x <= c * abs y |] ==> x = (#0::real)";
+Goal "[| c <= Numeral0; abs x <= c * abs y |] ==> x = (Numeral0::real)";
by (dtac order_le_imp_less_or_eq 1);
by Auto_tac;
-by (subgoal_tac "#0 <= c * abs y" 1);
+by (subgoal_tac "Numeral0 <= c * abs y" 1);
by (arith_tac 2);
by (asm_full_simp_tac (simpset() addsimps [real_0_le_mult_iff]) 1);
qed "rabs_ratiotest_lemma";
@@ -546,19 +546,19 @@
by (auto_tac (claset(),simpset() addsimps [le_Suc_ex]));
qed "le_Suc_ex_iff";
-(*All this trouble just to get #0<c *)
+(*All this trouble just to get Numeral0<c *)
Goal "[| ALL n. N <= n --> abs(f(Suc n)) <= c*abs(f n) |] \
-\ ==> #0 < c | summable f";
+\ ==> Numeral0 < c | summable f";
by (simp_tac (simpset() addsimps [linorder_not_le RS sym]) 1);
by (asm_full_simp_tac (simpset() addsimps [summable_Cauchy]) 1);
-by (Step_tac 1 THEN subgoal_tac "ALL n. N <= n --> f (Suc n) = #0" 1);
+by (Step_tac 1 THEN subgoal_tac "ALL n. N <= n --> f (Suc n) = Numeral0" 1);
by (blast_tac (claset() addIs [rabs_ratiotest_lemma]) 2);
by (res_inst_tac [("x","Suc N")] exI 1);
by (Clarify_tac 1);
by (dtac Suc_le_imp_diff_ge2 1 THEN Auto_tac);
qed "ratio_test_lemma2";
-Goal "[| c < #1; ALL n. N <= n --> abs(f(Suc n)) <= c*abs(f n) |] \
+Goal "[| c < Numeral1; ALL n. N <= n --> abs(f(Suc n)) <= c*abs(f n) |] \
\ ==> summable f";
by (forward_tac [ratio_test_lemma2] 1);
by Auto_tac;
@@ -573,7 +573,7 @@
by (auto_tac (claset() addIs [real_mult_le_mono1],
simpset() addsimps [summable_def]));
by (asm_full_simp_tac (simpset() addsimps real_mult_ac) 1);
-by (res_inst_tac [("x","abs(f N) * (#1/(#1 - c)) / (c ^ N)")] exI 1);
+by (res_inst_tac [("x","abs(f N) * (Numeral1/(Numeral1 - c)) / (c ^ N)")] exI 1);
by (rtac sums_divide 1);
by (rtac sums_mult 1);
by (auto_tac (claset() addSIs [sums_mult,geometric_sums],
--- a/src/HOL/Hyperreal/Series.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Hyperreal/Series.thy Fri Oct 05 21:52:39 2001 +0200
@@ -9,8 +9,8 @@
consts sumr :: "[nat,nat,(nat=>real)] => real"
primrec
- sumr_0 "sumr m 0 f = #0"
- sumr_Suc "sumr m (Suc n) f = (if n < m then #0
+ sumr_0 "sumr m 0 f = Numeral0"
+ sumr_Suc "sumr m (Suc n) f = (if n < m then Numeral0
else sumr m n f + f(n))"
constdefs
--- a/src/HOL/Hyperreal/hypreal_arith0.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Hyperreal/hypreal_arith0.ML Fri Oct 05 21:52:39 2001 +0200
@@ -115,7 +115,7 @@
qed "";
Goal "!!a::hypreal. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
-\ ==> #6*a <= #5*l+i";
+\ ==> # 6*a <= # 5*l+i";
by (fast_arith_tac 1);
qed "";
*)
--- a/src/HOL/IMP/Compiler.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/IMP/Compiler.thy Fri Oct 05 21:52:39 2001 +0200
@@ -39,9 +39,9 @@
"compile (x:==a) = [ASIN x a]"
"compile (c1;c2) = compile c1 @ compile c2"
"compile (IF b THEN c1 ELSE c2) =
- [JMPF b (length(compile c1)+2)] @ compile c1 @
+ [JMPF b (length(compile c1) + # 2)] @ compile c1 @
[JMPF (%x. False) (length(compile c2)+1)] @ compile c2"
-"compile (WHILE b DO c) = [JMPF b (length(compile c)+2)] @ compile c @
+"compile (WHILE b DO c) = [JMPF b (length(compile c) + # 2)] @ compile c @
[JMPB (length(compile c)+1)]"
declare nth_append[simp];
--- a/src/HOL/IMP/Examples.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/IMP/Examples.ML Fri Oct 05 21:52:39 2001 +0200
@@ -34,7 +34,7 @@
val step = resolve_tac evalc.intrs 1;
val simp = Asm_simp_tac 1;
Goalw [factorial_def] "a~=b ==> \
-\ <factorial a b, Mem(a:=#3)> -c-> Mem(b:=#6,a:=#0)";
+\ <factorial a b, Mem(a:=# 3)> -c-> Mem(b:=# 6,a:=Numeral0)";
by (ftac not_sym 1);
by step;
by step;
--- a/src/HOL/IMPP/EvenOdd.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/IMPP/EvenOdd.ML Fri Oct 05 21:52:39 2001 +0200
@@ -11,13 +11,13 @@
qed "even_0";
Addsimps [even_0];
-Goalw [even_def] "even 1' = False";
+Goalw [even_def] "even (Suc 0) = False";
by (Simp_tac 1);
qed "not_even_1";
Addsimps [not_even_1];
Goalw [even_def] "even (Suc (Suc n)) = even n";
-by (subgoal_tac "Suc (Suc n) = n+#2" 1);
+by (subgoal_tac "Suc (Suc n) = n+# 2" 1);
by (Simp_tac 2);
by (etac ssubst 1);
by (rtac dvd_reduce 1);
@@ -50,13 +50,13 @@
section "verification";
-Goalw [odd_def] "{{Z=Arg+0}. BODY Even .{Res_ok}}|-{Z=Arg+1'}. odd .{Res_ok}";
+Goalw [odd_def] "{{Z=Arg+0}. BODY Even .{Res_ok}}|-{Z=Arg+Suc 0}. odd .{Res_ok}";
by (rtac hoare_derivs.If 1);
by (rtac (hoare_derivs.Ass RS conseq1) 1);
by (clarsimp_tac Arg_Res_css 1);
by (rtac export_s 1);
by (rtac (hoare_derivs.Call RS conseq1) 1);
-by (res_inst_tac [("P","Z=Arg+2")] conseq12 1);
+by (res_inst_tac [("P","Z=Arg+Suc (Suc 0)")] conseq12 1);
by (rtac single_asm 1);
by (auto_tac Arg_Res_css);
qed "Odd_lemma";
--- a/src/HOL/IMPP/EvenOdd.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/IMPP/EvenOdd.thy Fri Oct 05 21:52:39 2001 +0200
@@ -9,7 +9,7 @@
EvenOdd = Misc +
constdefs even :: nat => bool
- "even n == #2 dvd n"
+ "even n == # 2 dvd n"
consts
Even, Odd :: pname
@@ -27,7 +27,7 @@
odd :: com
"odd == IF (%s. s<Arg>=0)
THEN Loc Res:==(%s. 1)
- ELSE(Loc Res:=CALL Even (%s. s<Arg> -1))"
+ ELSE(Loc Res:=CALL Even (%s. s<Arg> - 1))"
defs
bodies_def "bodies == [(Even,evn),(Odd,odd)]"
--- a/src/HOL/Induct/Com.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Induct/Com.thy Fri Oct 05 21:52:39 2001 +0200
@@ -52,10 +52,10 @@
IfTrue "[| (e,s) -|[eval]-> (0,s'); (c0,s') -[eval]-> s1 |]
==> (IF e THEN c0 ELSE c1, s) -[eval]-> s1"
- IfFalse "[| (e,s) -|[eval]-> (1',s'); (c1,s') -[eval]-> s1 |]
+ IfFalse "[| (e,s) -|[eval]-> (Suc 0, s'); (c1,s') -[eval]-> s1 |]
==> (IF e THEN c0 ELSE c1, s) -[eval]-> s1"
- WhileFalse "(e,s) -|[eval]-> (1',s1) ==> (WHILE e DO c, s) -[eval]-> s1"
+ WhileFalse "(e,s) -|[eval]-> (Suc 0, s1) ==> (WHILE e DO c, s) -[eval]-> s1"
WhileTrue "[| (e,s) -|[eval]-> (0,s1);
(c,s1) -[eval]-> s2; (WHILE e DO c, s2) -[eval]-> s3 |]
--- a/src/HOL/Induct/Mutil.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Induct/Mutil.thy Fri Oct 05 21:52:39 2001 +0200
@@ -29,7 +29,7 @@
constdefs
coloured :: "nat => (nat \<times> nat) set"
- "coloured b == {(i, j). (i + j) mod #2 = b}"
+ "coloured b == {(i, j). (i + j) mod # 2 = b}"
text {* \medskip The union of two disjoint tilings is a tiling *}
@@ -61,14 +61,14 @@
apply auto
done
-lemma dominoes_tile_row [intro!]: "{i} \<times> lessThan (#2 * n) \<in> tiling domino"
+lemma dominoes_tile_row [intro!]: "{i} \<times> lessThan (# 2 * n) \<in> tiling domino"
apply (induct n)
apply (simp_all add: Un_assoc [symmetric])
apply (rule tiling.Un)
apply (auto simp add: sing_Times_lemma)
done
-lemma dominoes_tile_matrix: "(lessThan m) \<times> lessThan (#2 * n) \<in> tiling domino"
+lemma dominoes_tile_matrix: "(lessThan m) \<times> lessThan (# 2 * n) \<in> tiling domino"
apply (induct m)
apply auto
done
@@ -78,7 +78,7 @@
lemma coloured_insert [simp]:
"coloured b \<inter> (insert (i, j) t) =
- (if (i + j) mod #2 = b then insert (i, j) (coloured b \<inter> t)
+ (if (i + j) mod # 2 = b then insert (i, j) (coloured b \<inter> t)
else coloured b \<inter> t)"
apply (unfold coloured_def)
apply auto
@@ -110,7 +110,7 @@
Diff_Int_distrib [simp]
lemma tiling_domino_0_1:
- "t \<in> tiling domino ==> card (coloured 0 \<inter> t) = card (coloured 1' \<inter> t)"
+ "t \<in> tiling domino ==> card (coloured 0 \<inter> t) = card (coloured (Suc 0) \<inter> t)"
apply (erule tiling.induct)
apply (drule_tac [2] domino_singletons)
apply auto
@@ -125,13 +125,13 @@
theorem gen_mutil_not_tiling:
"t \<in> tiling domino ==>
- (i + j) mod #2 = 0 ==> (m + n) mod #2 = 0 ==>
+ (i + j) mod # 2 = 0 ==> (m + n) mod # 2 = 0 ==>
{(i, j), (m, n)} \<subseteq> t
==> (t - {(i, j)} - {(m, n)}) \<notin> tiling domino"
apply (rule notI)
apply (subgoal_tac
"card (coloured 0 \<inter> (t - {(i, j)} - {(m, n)})) <
- card (coloured 1' \<inter> (t - {(i, j)} - {(m, n)}))")
+ card (coloured (Suc 0) \<inter> (t - {(i, j)} - {(m, n)}))")
apply (force simp only: tiling_domino_0_1)
apply (simp add: tiling_domino_0_1 [symmetric])
apply (simp add: coloured_def card_Diff2_less)
@@ -140,8 +140,8 @@
text {* Apply the general theorem to the well-known case *}
theorem mutil_not_tiling:
- "t = lessThan (#2 * Suc m) \<times> lessThan (#2 * Suc n)
- ==> t - {(0, 0)} - {(Suc (#2 * m), Suc (#2 * n))} \<notin> tiling domino"
+ "t = lessThan (# 2 * Suc m) \<times> lessThan (# 2 * Suc n)
+ ==> t - {(0, 0)} - {(Suc (# 2 * m), Suc (# 2 * n))} \<notin> tiling domino"
apply (rule gen_mutil_not_tiling)
apply (blast intro!: dominoes_tile_matrix)
apply auto
--- a/src/HOL/Integ/Bin.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Integ/Bin.ML Fri Oct 05 21:52:39 2001 +0200
@@ -160,7 +160,7 @@
(*The correctness of shifting. But it doesn't seem to give a measurable
speed-up.*)
-Goal "(#2::int) * number_of w = number_of (w BIT False)";
+Goal "(# 2::int) * number_of w = number_of (w BIT False)";
by (induct_tac "w" 1);
by (ALLGOALS (asm_simp_tac
(simpset() addsimps bin_mult_simps @ [zadd_zmult_distrib] @ zadd_ac)));
@@ -169,11 +169,11 @@
(** Simplification rules with integer constants **)
-Goal "#0 + z = (z::int)";
+Goal "Numeral0 + z = (z::int)";
by (Simp_tac 1);
qed "zadd_0";
-Goal "z + #0 = (z::int)";
+Goal "z + Numeral0 = (z::int)";
by (Simp_tac 1);
qed "zadd_0_right";
@@ -182,29 +182,29 @@
(** Converting simple cases of (int n) to numerals **)
-(*int 0 = #0 *)
+(*int 0 = Numeral0 *)
bind_thm ("int_0", number_of_Pls RS sym);
-Goal "int (Suc n) = #1 + int n";
+Goal "int (Suc n) = Numeral1 + int n";
by (simp_tac (simpset() addsimps [zadd_int]) 1);
qed "int_Suc";
-Goal "- (#0) = (#0::int)";
+Goal "- (Numeral0) = (Numeral0::int)";
by (Simp_tac 1);
qed "zminus_0";
Addsimps [zminus_0];
-Goal "(#0::int) - x = -x";
+Goal "(Numeral0::int) - x = -x";
by (simp_tac (simpset() addsimps [zdiff_def]) 1);
qed "zdiff0";
-Goal "x - (#0::int) = x";
+Goal "x - (Numeral0::int) = x";
by (simp_tac (simpset() addsimps [zdiff_def]) 1);
qed "zdiff0_right";
-Goal "x - x = (#0::int)";
+Goal "x - x = (Numeral0::int)";
by (simp_tac (simpset() addsimps [zdiff_def]) 1);
qed "zdiff_self";
@@ -234,27 +234,27 @@
(** Special-case simplification for small constants **)
-Goal "#0 * z = (#0::int)";
+Goal "Numeral0 * z = (Numeral0::int)";
by (Simp_tac 1);
qed "zmult_0";
-Goal "z * #0 = (#0::int)";
+Goal "z * Numeral0 = (Numeral0::int)";
by (Simp_tac 1);
qed "zmult_0_right";
-Goal "#1 * z = (z::int)";
+Goal "Numeral1 * z = (z::int)";
by (Simp_tac 1);
qed "zmult_1";
-Goal "z * #1 = (z::int)";
+Goal "z * Numeral1 = (z::int)";
by (Simp_tac 1);
qed "zmult_1_right";
-Goal "#-1 * z = -(z::int)";
+Goal "# -1 * z = -(z::int)";
by (simp_tac (simpset() addsimps zcompare_rls@[zmult_zminus]) 1);
qed "zmult_minus1";
-Goal "z * #-1 = -(z::int)";
+Goal "z * # -1 = -(z::int)";
by (simp_tac (simpset() addsimps zcompare_rls@[zmult_zminus_right]) 1);
qed "zmult_minus1_right";
@@ -272,61 +272,61 @@
(** Inequality reasoning **)
-Goal "(m*n = (#0::int)) = (m = #0 | n = #0)";
+Goal "(m*n = (Numeral0::int)) = (m = Numeral0 | n = Numeral0)";
by (stac (int_0 RS sym) 1 THEN rtac zmult_eq_int0_iff 1);
qed "zmult_eq_0_iff";
AddIffs [zmult_eq_0_iff];
-Goal "(w < z + (#1::int)) = (w<z | w=z)";
+Goal "(w < z + (Numeral1::int)) = (w<z | w=z)";
by (simp_tac (simpset() addsimps [zless_add_int_Suc_eq]) 1);
qed "zless_add1_eq";
-Goal "(w + (#1::int) <= z) = (w<z)";
+Goal "(w + (Numeral1::int) <= z) = (w<z)";
by (simp_tac (simpset() addsimps [add_int_Suc_zle_eq]) 1);
qed "add1_zle_eq";
-Goal "((#1::int) + w <= z) = (w<z)";
+Goal "((Numeral1::int) + w <= z) = (w<z)";
by (stac zadd_commute 1);
by (rtac add1_zle_eq 1);
qed "add1_left_zle_eq";
-Goal "neg x = (x < #0)";
+Goal "neg x = (x < Numeral0)";
by (simp_tac (simpset() addsimps [neg_eq_less_int0]) 1);
qed "neg_eq_less_0";
-Goal "(~neg x) = (#0 <= x)";
+Goal "(~neg x) = (Numeral0 <= x)";
by (simp_tac (simpset() addsimps [not_neg_eq_ge_int0]) 1);
qed "not_neg_eq_ge_0";
-Goal "#0 <= int m";
+Goal "Numeral0 <= int m";
by (Simp_tac 1);
qed "zero_zle_int";
AddIffs [zero_zle_int];
-(** Needed because (int 0) rewrites to #0.
+(** Needed because (int 0) rewrites to Numeral0. (* FIXME !? *)
Can these be generalized without evaluating large numbers?**)
-Goal "~ (int k < #0)";
+Goal "~ (int k < Numeral0)";
by (Simp_tac 1);
qed "int_less_0_conv";
-Goal "(int k <= #0) = (k=0)";
+Goal "(int k <= Numeral0) = (k=0)";
by (Simp_tac 1);
qed "int_le_0_conv";
-Goal "(int k = #0) = (k=0)";
+Goal "(int k = Numeral0) = (k=0)";
by (Simp_tac 1);
qed "int_eq_0_conv";
-Goal "(#0 < int k) = (0<k)";
+Goal "(Numeral0 < int k) = (0<k)";
by (Simp_tac 1);
qed "zero_less_int_conv";
Addsimps [int_less_0_conv, int_le_0_conv, int_eq_0_conv, zero_less_int_conv];
-Goal "(0 < nat z) = (#0 < z)";
-by (cut_inst_tac [("w","#0")] zless_nat_conj 1);
+Goal "(0 < nat z) = (Numeral0 < z)";
+by (cut_inst_tac [("w","Numeral0")] zless_nat_conj 1);
by Auto_tac;
qed "zero_less_nat_eq";
Addsimps [zero_less_nat_eq];
@@ -339,7 +339,7 @@
Goalw [iszero_def]
"((number_of x::int) = number_of y) = \
\ iszero (number_of (bin_add x (bin_minus y)))";
-by (simp_tac (simpset() delsimps [number_of_reorient]
+by (simp_tac (simpset() delsimps [thm "number_of_reorient"]
addsimps zcompare_rls @ [number_of_add, number_of_minus]) 1);
qed "eq_number_of_eq";
--- a/src/HOL/Integ/Int.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Integ/Int.ML Fri Oct 05 21:52:39 2001 +0200
@@ -50,7 +50,7 @@
val eq_diff_eq = eq_zdiff_eq
val eqI_rules = [zless_eqI, zeq_eqI, zle_eqI]
fun dest_eqI th =
- #1 (HOLogic.dest_bin "op =" HOLogic.boolT
+ #1 (HOLogic.dest_bin "op =" HOLogic.boolT
(HOLogic.dest_Trueprop (concl_of th)))
val diff_def = zdiff_def
--- a/src/HOL/Integ/IntArith.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Integ/IntArith.ML Fri Oct 05 21:52:39 2001 +0200
@@ -20,7 +20,7 @@
(*** Intermediate value theorems ***)
-Goal "(ALL i<n. abs(f(i+1) - f i) <= #1) --> \
+Goal "(ALL i<n::nat. abs(f(i+1) - f i) <= Numeral1) --> \
\ f 0 <= k --> k <= f n --> (EX i <= n. f i = (k::int))";
by(induct_tac "n" 1);
by(Asm_simp_tac 1);
@@ -40,7 +40,7 @@
bind_thm("nat0_intermed_int_val", rulify_no_asm lemma);
-Goal "[| !i. m <= i & i < n --> abs(f(i+1) - f i) <= #1; m < n; \
+Goal "[| !i. m <= i & i < n --> abs(f(i + 1::nat) - f i) <= Numeral1; m < n; \
\ f m <= k; k <= f n |] ==> ? i. m <= i & i <= n & f i = (k::int)";
by(cut_inst_tac [("n","n-m"),("f", "%i. f(i+m)"),("k","k")]lemma 1);
by(Asm_full_simp_tac 1);
@@ -56,22 +56,22 @@
(*** Some convenient biconditionals for products of signs ***)
-Goal "[| (#0::int) < i; #0 < j |] ==> #0 < i*j";
+Goal "[| (Numeral0::int) < i; Numeral0 < j |] ==> Numeral0 < i*j";
by (dtac zmult_zless_mono1 1);
by Auto_tac;
qed "zmult_pos";
-Goal "[| i < (#0::int); j < #0 |] ==> #0 < i*j";
+Goal "[| i < (Numeral0::int); j < Numeral0 |] ==> Numeral0 < i*j";
by (dtac zmult_zless_mono1_neg 1);
by Auto_tac;
qed "zmult_neg";
-Goal "[| (#0::int) < i; j < #0 |] ==> i*j < #0";
+Goal "[| (Numeral0::int) < i; j < Numeral0 |] ==> i*j < Numeral0";
by (dtac zmult_zless_mono1_neg 1);
by Auto_tac;
qed "zmult_pos_neg";
-Goal "((#0::int) < x*y) = (#0 < x & #0 < y | x < #0 & y < #0)";
+Goal "((Numeral0::int) < x*y) = (Numeral0 < x & Numeral0 < y | x < Numeral0 & y < Numeral0)";
by (auto_tac (claset(),
simpset() addsimps [order_le_less, linorder_not_less,
zmult_pos, zmult_neg]));
@@ -84,13 +84,13 @@
simpset() addsimps [zmult_commute]));
qed "int_0_less_mult_iff";
-Goal "((#0::int) <= x*y) = (#0 <= x & #0 <= y | x <= #0 & y <= #0)";
+Goal "((Numeral0::int) <= x*y) = (Numeral0 <= x & Numeral0 <= y | x <= Numeral0 & y <= Numeral0)";
by (auto_tac (claset(),
simpset() addsimps [order_le_less, linorder_not_less,
int_0_less_mult_iff]));
qed "int_0_le_mult_iff";
-Goal "(x*y < (#0::int)) = (#0 < x & y < #0 | x < #0 & #0 < y)";
+Goal "(x*y < (Numeral0::int)) = (Numeral0 < x & y < Numeral0 | x < Numeral0 & Numeral0 < y)";
by (auto_tac (claset(),
simpset() addsimps [int_0_le_mult_iff,
linorder_not_le RS sym]));
@@ -98,7 +98,7 @@
simpset() addsimps [linorder_not_le]));
qed "zmult_less_0_iff";
-Goal "(x*y <= (#0::int)) = (#0 <= x & y <= #0 | x <= #0 & #0 <= y)";
+Goal "(x*y <= (Numeral0::int)) = (Numeral0 <= x & y <= Numeral0 | x <= Numeral0 & Numeral0 <= y)";
by (auto_tac (claset() addDs [order_less_not_sym],
simpset() addsimps [int_0_less_mult_iff,
linorder_not_less RS sym]));
@@ -109,19 +109,19 @@
addsimps [zmult_less_0_iff, zle_def]) 1);
qed "abs_mult";
-Goal "(abs x = #0) = (x = (#0::int))";
+Goal "(abs x = Numeral0) = (x = (Numeral0::int))";
by (simp_tac (simpset () addsplits [zabs_split]) 1);
qed "abs_eq_0";
AddIffs [abs_eq_0];
-Goal "(#0 < abs x) = (x ~= (#0::int))";
+Goal "(Numeral0 < abs x) = (x ~= (Numeral0::int))";
by (simp_tac (simpset () addsplits [zabs_split]) 1);
by (arith_tac 1);
qed "zero_less_abs_iff";
AddIffs [zero_less_abs_iff];
-Goal "#0 <= x * (x::int)";
-by (subgoal_tac "(- x) * x <= #0" 1);
+Goal "Numeral0 <= x * (x::int)";
+by (subgoal_tac "(- x) * x <= Numeral0" 1);
by (Asm_full_simp_tac 1);
by (simp_tac (HOL_basic_ss addsimps [zmult_le_0_iff]) 1);
by Auto_tac;
@@ -132,48 +132,48 @@
(*** Products and 1, by T. M. Rasmussen ***)
-Goal "(m = m*(n::int)) = (n = #1 | m = #0)";
+Goal "(m = m*(n::int)) = (n = Numeral1 | m = Numeral0)";
by Auto_tac;
-by (subgoal_tac "m*#1 = m*n" 1);
+by (subgoal_tac "m*Numeral1 = m*n" 1);
by (dtac (zmult_cancel1 RS iffD1) 1);
by Auto_tac;
qed "zmult_eq_self_iff";
-Goal "[| #1 < m; #1 < n |] ==> #1 < m*(n::int)";
-by (res_inst_tac [("y","#1*n")] order_less_trans 1);
+Goal "[| Numeral1 < m; Numeral1 < n |] ==> Numeral1 < m*(n::int)";
+by (res_inst_tac [("y","Numeral1*n")] order_less_trans 1);
by (rtac zmult_zless_mono1 2);
by (ALLGOALS Asm_simp_tac);
qed "zless_1_zmult";
-Goal "[| #0 < n; n ~= #1 |] ==> #1 < (n::int)";
+Goal "[| Numeral0 < n; n ~= Numeral1 |] ==> Numeral1 < (n::int)";
by (arith_tac 1);
val lemma = result();
-Goal "#0 < (m::int) ==> (m * n = #1) = (m = #1 & n = #1)";
+Goal "Numeral0 < (m::int) ==> (m * n = Numeral1) = (m = Numeral1 & n = Numeral1)";
by Auto_tac;
-by (case_tac "m=#1" 1);
-by (case_tac "n=#1" 2);
-by (case_tac "m=#1" 4);
-by (case_tac "n=#1" 5);
+by (case_tac "m=Numeral1" 1);
+by (case_tac "n=Numeral1" 2);
+by (case_tac "m=Numeral1" 4);
+by (case_tac "n=Numeral1" 5);
by Auto_tac;
by distinct_subgoals_tac;
-by (subgoal_tac "#1<m*n" 1);
+by (subgoal_tac "Numeral1<m*n" 1);
by (Asm_full_simp_tac 1);
by (rtac zless_1_zmult 1);
by (ALLGOALS (rtac lemma));
by Auto_tac;
-by (subgoal_tac "#0<m*n" 1);
+by (subgoal_tac "Numeral0<m*n" 1);
by (Asm_simp_tac 2);
by (dtac (int_0_less_mult_iff RS iffD1) 1);
by Auto_tac;
qed "pos_zmult_eq_1_iff";
-Goal "(m*n = (#1::int)) = ((m = #1 & n = #1) | (m = #-1 & n = #-1))";
-by (case_tac "#0<m" 1);
+Goal "(m*n = (Numeral1::int)) = ((m = Numeral1 & n = Numeral1) | (m = # -1 & n = # -1))";
+by (case_tac "Numeral0<m" 1);
by (asm_simp_tac (simpset() addsimps [pos_zmult_eq_1_iff]) 1);
-by (case_tac "m=#0" 1);
+by (case_tac "m=Numeral0" 1);
by (Asm_simp_tac 1);
-by (subgoal_tac "#0 < -m" 1);
+by (subgoal_tac "Numeral0 < -m" 1);
by (arith_tac 2);
by (dres_inst_tac [("n","-n")] pos_zmult_eq_1_iff 1);
by Auto_tac;
--- a/src/HOL/Integ/IntDef.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Integ/IntDef.ML Fri Oct 05 21:52:39 2001 +0200
@@ -7,8 +7,8 @@
*)
-(*Rewrite the overloaded 0::int to (int 0)*)
-Addsimps [Zero_def];
+(*Rewrite the overloaded 0::int to (int 0)*) (* FIXME !? *)
+Addsimps [Zero_int_def];
Goalw [intrel_def] "(((x1,y1),(x2,y2)): intrel) = (x1+y2 = x2+y1)";
by (Blast_tac 1);
@@ -326,7 +326,7 @@
by (asm_simp_tac (simpset() addsimps [zmult]) 1);
qed "zmult_int0";
-Goalw [int_def] "int 1' * z = z";
+Goalw [int_def] "int (Suc 0) * z = z";
by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
by (asm_simp_tac (simpset() addsimps [zmult]) 1);
qed "zmult_int1";
@@ -335,7 +335,7 @@
by (rtac ([zmult_commute, zmult_int0] MRS trans) 1);
qed "zmult_int0_right";
-Goal "z * int 1' = z";
+Goal "z * int (Suc 0) = z";
by (rtac ([zmult_commute, zmult_int1] MRS trans) 1);
qed "zmult_int1_right";
--- a/src/HOL/Integ/IntDef.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Integ/IntDef.thy Fri Oct 05 21:52:39 2001 +0200
@@ -35,7 +35,7 @@
defs (*of overloaded constants*)
- Zero_def "0 == int 0"
+ Zero_int_def "0 == int 0"
zadd_def
"z + w ==
--- a/src/HOL/Integ/IntDiv.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Integ/IntDiv.ML Fri Oct 05 21:52:39 2001 +0200
@@ -34,21 +34,21 @@
(*** Uniqueness and monotonicity of quotients and remainders ***)
-Goal "[| b*q' + r' <= b*q + r; #0 <= r'; #0 < b; r < b |] \
+Goal "[| b*q' + r' <= b*q + r; Numeral0 <= r'; Numeral0 < b; r < b |] \
\ ==> q' <= (q::int)";
by (subgoal_tac "r' + b * (q'-q) <= r" 1);
by (simp_tac (simpset() addsimps [zdiff_zmult_distrib2]) 2);
-by (subgoal_tac "#0 < b * (#1 + q - q')" 1);
+by (subgoal_tac "Numeral0 < b * (Numeral1 + q - q')" 1);
by (etac order_le_less_trans 2);
by (full_simp_tac (simpset() addsimps [zdiff_zmult_distrib2,
zadd_zmult_distrib2]) 2);
-by (subgoal_tac "b * q' < b * (#1 + q)" 1);
+by (subgoal_tac "b * q' < b * (Numeral1 + q)" 1);
by (full_simp_tac (simpset() addsimps [zdiff_zmult_distrib2,
zadd_zmult_distrib2]) 2);
by (asm_full_simp_tac (simpset() addsimps [zmult_zless_cancel1]) 1);
qed "unique_quotient_lemma";
-Goal "[| b*q' + r' <= b*q + r; r <= #0; b < #0; b < r' |] \
+Goal "[| b*q' + r' <= b*q + r; r <= Numeral0; b < Numeral0; b < r' |] \
\ ==> q <= (q'::int)";
by (res_inst_tac [("b", "-b"), ("r", "-r'"), ("r'", "-r")]
unique_quotient_lemma 1);
@@ -57,7 +57,7 @@
qed "unique_quotient_lemma_neg";
-Goal "[| quorem ((a,b), (q,r)); quorem ((a,b), (q',r')); b ~= #0 |] \
+Goal "[| quorem ((a,b), (q,r)); quorem ((a,b), (q',r')); b ~= Numeral0 |] \
\ ==> q = q'";
by (asm_full_simp_tac
(simpset() addsimps split_ifs@
@@ -72,7 +72,7 @@
qed "unique_quotient";
-Goal "[| quorem ((a,b), (q,r)); quorem ((a,b), (q',r')); b ~= #0 |] \
+Goal "[| quorem ((a,b), (q,r)); quorem ((a,b), (q',r')); b ~= Numeral0 |] \
\ ==> r = r'";
by (subgoal_tac "q = q'" 1);
by (blast_tac (claset() addIs [unique_quotient]) 2);
@@ -84,8 +84,8 @@
Goal "adjust a b (q,r) = (let diff = r-b in \
-\ if #0 <= diff then (#2*q + #1, diff) \
-\ else (#2*q, r))";
+\ if Numeral0 <= diff then (# 2*q + Numeral1, diff) \
+\ else (# 2*q, r))";
by (simp_tac (simpset() addsimps [Let_def,adjust_def]) 1);
qed "adjust_eq";
Addsimps [adjust_eq];
@@ -101,9 +101,9 @@
bind_thm ("posDivAlg_raw_eqn", lemma RS hd posDivAlg.simps);
(**use with simproc to avoid re-proving the premise*)
-Goal "#0 < b ==> \
+Goal "Numeral0 < b ==> \
\ posDivAlg (a,b) = \
-\ (if a<b then (#0,a) else adjust a b (posDivAlg(a, #2*b)))";
+\ (if a<b then (Numeral0,a) else adjust a b (posDivAlg(a, # 2*b)))";
by (rtac (posDivAlg_raw_eqn RS trans) 1);
by (Asm_simp_tac 1);
qed "posDivAlg_eqn";
@@ -112,7 +112,7 @@
(*Correctness of posDivAlg: it computes quotients correctly*)
-Goal "#0 <= a --> #0 < b --> quorem ((a, b), posDivAlg (a, b))";
+Goal "Numeral0 <= a --> Numeral0 < b --> quorem ((a, b), posDivAlg (a, b))";
by (induct_thm_tac posDivAlg_induct "a b" 1);
by Auto_tac;
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [quorem_def])));
@@ -139,9 +139,9 @@
bind_thm ("negDivAlg_raw_eqn", lemma RS hd negDivAlg.simps);
(**use with simproc to avoid re-proving the premise*)
-Goal "#0 < b ==> \
+Goal "Numeral0 < b ==> \
\ negDivAlg (a,b) = \
-\ (if #0<=a+b then (#-1,a+b) else adjust a b (negDivAlg(a, #2*b)))";
+\ (if Numeral0<=a+b then (# -1,a+b) else adjust a b (negDivAlg(a, # 2*b)))";
by (rtac (negDivAlg_raw_eqn RS trans) 1);
by (Asm_simp_tac 1);
qed "negDivAlg_eqn";
@@ -151,7 +151,7 @@
(*Correctness of negDivAlg: it computes quotients correctly
It doesn't work if a=0 because the 0/b=0 rather than -1*)
-Goal "a < #0 --> #0 < b --> quorem ((a, b), negDivAlg (a, b))";
+Goal "a < Numeral0 --> Numeral0 < b --> quorem ((a, b), negDivAlg (a, b))";
by (induct_thm_tac negDivAlg_induct "a b" 1);
by Auto_tac;
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [quorem_def])));
@@ -168,18 +168,18 @@
(*** Existence shown by proving the division algorithm to be correct ***)
(*the case a=0*)
-Goal "b ~= #0 ==> quorem ((#0,b), (#0,#0))";
+Goal "b ~= Numeral0 ==> quorem ((Numeral0,b), (Numeral0,Numeral0))";
by (auto_tac (claset(),
simpset() addsimps [quorem_def, linorder_neq_iff]));
qed "quorem_0";
-Goal "posDivAlg (#0, b) = (#0, #0)";
+Goal "posDivAlg (Numeral0, b) = (Numeral0, Numeral0)";
by (stac posDivAlg_raw_eqn 1);
by Auto_tac;
qed "posDivAlg_0";
Addsimps [posDivAlg_0];
-Goal "negDivAlg (#-1, b) = (#-1, b-#1)";
+Goal "negDivAlg (# -1, b) = (# -1, b-Numeral1)";
by (stac negDivAlg_raw_eqn 1);
by Auto_tac;
qed "negDivAlg_minus1";
@@ -194,7 +194,7 @@
by (auto_tac (claset(), simpset() addsimps split_ifs@[quorem_def]));
qed "quorem_neg";
-Goal "b ~= #0 ==> quorem ((a,b), divAlg(a,b))";
+Goal "b ~= Numeral0 ==> quorem ((a,b), divAlg(a,b))";
by (auto_tac (claset(),
simpset() addsimps [quorem_0, divAlg_def]));
by (REPEAT_FIRST (resolve_tac [quorem_neg, posDivAlg_correct,
@@ -206,11 +206,11 @@
(** Arbitrary definitions for division by zero. Useful to simplify
certain equations **)
-Goal "a div (#0::int) = #0";
+Goal "a div (Numeral0::int) = Numeral0";
by (simp_tac (simpset() addsimps [div_def, divAlg_def, posDivAlg_raw_eqn]) 1);
qed "DIVISION_BY_ZERO_ZDIV"; (*NOT for adding to default simpset*)
-Goal "a mod (#0::int) = a";
+Goal "a mod (Numeral0::int) = a";
by (simp_tac (simpset() addsimps [mod_def, divAlg_def, posDivAlg_raw_eqn]) 1);
qed "DIVISION_BY_ZERO_ZMOD"; (*NOT for adding to default simpset*)
@@ -222,20 +222,20 @@
(** Basic laws about division and remainder **)
Goal "(a::int) = b * (a div b) + (a mod b)";
-by (zdiv_undefined_case_tac "b = #0" 1);
+by (zdiv_undefined_case_tac "b = Numeral0" 1);
by (cut_inst_tac [("a","a"),("b","b")] divAlg_correct 1);
by (auto_tac (claset(),
simpset() addsimps [quorem_def, div_def, mod_def]));
qed "zmod_zdiv_equality";
-Goal "(#0::int) < b ==> #0 <= a mod b & a mod b < b";
+Goal "(Numeral0::int) < b ==> Numeral0 <= a mod b & a mod b < b";
by (cut_inst_tac [("a","a"),("b","b")] divAlg_correct 1);
by (auto_tac (claset(),
simpset() addsimps [quorem_def, mod_def]));
bind_thm ("pos_mod_sign", result() RS conjunct1);
bind_thm ("pos_mod_bound", result() RS conjunct2);
-Goal "b < (#0::int) ==> a mod b <= #0 & b < a mod b";
+Goal "b < (Numeral0::int) ==> a mod b <= Numeral0 & b < a mod b";
by (cut_inst_tac [("a","a"),("b","b")] divAlg_correct 1);
by (auto_tac (claset(),
simpset() addsimps [quorem_def, div_def, mod_def]));
@@ -245,7 +245,7 @@
(** proving general properties of div and mod **)
-Goal "b ~= #0 ==> quorem ((a, b), (a div b, a mod b))";
+Goal "b ~= Numeral0 ==> quorem ((a, b), (a div b, a mod b))";
by (cut_inst_tac [("a","a"),("b","b")] zmod_zdiv_equality 1);
by (auto_tac
(claset(),
@@ -254,43 +254,43 @@
neg_mod_sign, neg_mod_bound]));
qed "quorem_div_mod";
-Goal "[| quorem((a,b),(q,r)); b ~= #0 |] ==> a div b = q";
+Goal "[| quorem((a,b),(q,r)); b ~= Numeral0 |] ==> a div b = q";
by (asm_simp_tac (simpset() addsimps [quorem_div_mod RS unique_quotient]) 1);
qed "quorem_div";
-Goal "[| quorem((a,b),(q,r)); b ~= #0 |] ==> a mod b = r";
+Goal "[| quorem((a,b),(q,r)); b ~= Numeral0 |] ==> a mod b = r";
by (asm_simp_tac (simpset() addsimps [quorem_div_mod RS unique_remainder]) 1);
qed "quorem_mod";
-Goal "[| (#0::int) <= a; a < b |] ==> a div b = #0";
+Goal "[| (Numeral0::int) <= a; a < b |] ==> a div b = Numeral0";
by (rtac quorem_div 1);
by (auto_tac (claset(), simpset() addsimps [quorem_def]));
qed "div_pos_pos_trivial";
-Goal "[| a <= (#0::int); b < a |] ==> a div b = #0";
+Goal "[| a <= (Numeral0::int); b < a |] ==> a div b = Numeral0";
by (rtac quorem_div 1);
by (auto_tac (claset(), simpset() addsimps [quorem_def]));
qed "div_neg_neg_trivial";
-Goal "[| (#0::int) < a; a+b <= #0 |] ==> a div b = #-1";
+Goal "[| (Numeral0::int) < a; a+b <= Numeral0 |] ==> a div b = # -1";
by (rtac quorem_div 1);
by (auto_tac (claset(), simpset() addsimps [quorem_def]));
qed "div_pos_neg_trivial";
-(*There is no div_neg_pos_trivial because #0 div b = #0 would supersede it*)
+(*There is no div_neg_pos_trivial because Numeral0 div b = Numeral0 would supersede it*)
-Goal "[| (#0::int) <= a; a < b |] ==> a mod b = a";
-by (res_inst_tac [("q","#0")] quorem_mod 1);
+Goal "[| (Numeral0::int) <= a; a < b |] ==> a mod b = a";
+by (res_inst_tac [("q","Numeral0")] quorem_mod 1);
by (auto_tac (claset(), simpset() addsimps [quorem_def]));
qed "mod_pos_pos_trivial";
-Goal "[| a <= (#0::int); b < a |] ==> a mod b = a";
-by (res_inst_tac [("q","#0")] quorem_mod 1);
+Goal "[| a <= (Numeral0::int); b < a |] ==> a mod b = a";
+by (res_inst_tac [("q","Numeral0")] quorem_mod 1);
by (auto_tac (claset(), simpset() addsimps [quorem_def]));
qed "mod_neg_neg_trivial";
-Goal "[| (#0::int) < a; a+b <= #0 |] ==> a mod b = a+b";
-by (res_inst_tac [("q","#-1")] quorem_mod 1);
+Goal "[| (Numeral0::int) < a; a+b <= Numeral0 |] ==> a mod b = a+b";
+by (res_inst_tac [("q","# -1")] quorem_mod 1);
by (auto_tac (claset(), simpset() addsimps [quorem_def]));
qed "mod_pos_neg_trivial";
@@ -299,7 +299,7 @@
(*Simpler laws such as -a div b = -(a div b) FAIL, but see just below*)
Goal "(-a) div (-b) = a div (b::int)";
-by (zdiv_undefined_case_tac "b = #0" 1);
+by (zdiv_undefined_case_tac "b = Numeral0" 1);
by (stac ((simplify(simpset()) (quorem_div_mod RS quorem_neg))
RS quorem_div) 1);
by Auto_tac;
@@ -308,7 +308,7 @@
(*Simpler laws such as -a mod b = -(a mod b) FAIL, but see just below*)
Goal "(-a) mod (-b) = - (a mod (b::int))";
-by (zdiv_undefined_case_tac "b = #0" 1);
+by (zdiv_undefined_case_tac "b = Numeral0" 1);
by (stac ((simplify(simpset()) (quorem_div_mod RS quorem_neg))
RS quorem_mod) 1);
by Auto_tac;
@@ -319,8 +319,8 @@
(*** div, mod and unary minus ***)
Goal "quorem((a,b),(q,r)) \
-\ ==> quorem ((-a,b), (if r=#0 then -q else -q-#1), \
-\ (if r=#0 then #0 else b-r))";
+\ ==> quorem ((-a,b), (if r=Numeral0 then -q else -q-Numeral1), \
+\ (if r=Numeral0 then Numeral0 else b-r))";
by (auto_tac
(claset(),
simpset() addsimps split_ifs@
@@ -328,14 +328,14 @@
zdiff_zmult_distrib2]));
val lemma = result();
-Goal "b ~= (#0::int) \
+Goal "b ~= (Numeral0::int) \
\ ==> (-a) div b = \
-\ (if a mod b = #0 then - (a div b) else - (a div b) - #1)";
+\ (if a mod b = Numeral0 then - (a div b) else - (a div b) - Numeral1)";
by (blast_tac (claset() addIs [quorem_div_mod RS lemma RS quorem_div]) 1);
qed "zdiv_zminus1_eq_if";
-Goal "(-a::int) mod b = (if a mod b = #0 then #0 else b - (a mod b))";
-by (zdiv_undefined_case_tac "b = #0" 1);
+Goal "(-a::int) mod b = (if a mod b = Numeral0 then Numeral0 else b - (a mod b))";
+by (zdiv_undefined_case_tac "b = Numeral0" 1);
by (blast_tac (claset() addIs [quorem_div_mod RS lemma RS quorem_mod]) 1);
qed "zmod_zminus1_eq_if";
@@ -349,32 +349,32 @@
by Auto_tac;
qed "zmod_zminus2";
-Goal "b ~= (#0::int) \
+Goal "b ~= (Numeral0::int) \
\ ==> a div (-b) = \
-\ (if a mod b = #0 then - (a div b) else - (a div b) - #1)";
+\ (if a mod b = Numeral0 then - (a div b) else - (a div b) - Numeral1)";
by (asm_simp_tac (simpset() addsimps [zdiv_zminus1_eq_if, zdiv_zminus2]) 1);
qed "zdiv_zminus2_eq_if";
-Goal "a mod (-b::int) = (if a mod b = #0 then #0 else (a mod b) - b)";
+Goal "a mod (-b::int) = (if a mod b = Numeral0 then Numeral0 else (a mod b) - b)";
by (asm_simp_tac (simpset() addsimps [zmod_zminus1_eq_if, zmod_zminus2]) 1);
qed "zmod_zminus2_eq_if";
(*** division of a number by itself ***)
-Goal "[| (#0::int) < a; a = r + a*q; r < a |] ==> #1 <= q";
-by (subgoal_tac "#0 < a*q" 1);
+Goal "[| (Numeral0::int) < a; a = r + a*q; r < a |] ==> Numeral1 <= q";
+by (subgoal_tac "Numeral0 < a*q" 1);
by (arith_tac 2);
by (asm_full_simp_tac (simpset() addsimps [int_0_less_mult_iff]) 1);
val lemma1 = result();
-Goal "[| (#0::int) < a; a = r + a*q; #0 <= r |] ==> q <= #1";
-by (subgoal_tac "#0 <= a*(#1-q)" 1);
+Goal "[| (Numeral0::int) < a; a = r + a*q; Numeral0 <= r |] ==> q <= Numeral1";
+by (subgoal_tac "Numeral0 <= a*(Numeral1-q)" 1);
by (asm_simp_tac (simpset() addsimps [zdiff_zmult_distrib2]) 2);
by (asm_full_simp_tac (simpset() addsimps [int_0_le_mult_iff]) 1);
val lemma2 = result();
-Goal "[| quorem((a,a),(q,r)); a ~= (#0::int) |] ==> q = #1";
+Goal "[| quorem((a,a),(q,r)); a ~= (Numeral0::int) |] ==> q = Numeral1";
by (asm_full_simp_tac
(simpset() addsimps split_ifs@[quorem_def, linorder_neq_iff]) 1);
by (rtac order_antisym 1);
@@ -386,20 +386,20 @@
simpset() addsimps [zadd_commute, zmult_zminus]) 1));
qed "self_quotient";
-Goal "[| quorem((a,a),(q,r)); a ~= (#0::int) |] ==> r = #0";
+Goal "[| quorem((a,a),(q,r)); a ~= (Numeral0::int) |] ==> r = Numeral0";
by (ftac self_quotient 1);
by (assume_tac 1);
by (asm_full_simp_tac (simpset() addsimps [quorem_def]) 1);
qed "self_remainder";
-Goal "a ~= #0 ==> a div a = (#1::int)";
+Goal "a ~= Numeral0 ==> a div a = (Numeral1::int)";
by (asm_simp_tac (simpset() addsimps [quorem_div_mod RS self_quotient]) 1);
qed "zdiv_self";
Addsimps [zdiv_self];
(*Here we have 0 mod 0 = 0, also assumed by Knuth (who puts m mod 0 = 0) *)
-Goal "a mod a = (#0::int)";
-by (zdiv_undefined_case_tac "a = #0" 1);
+Goal "a mod a = (Numeral0::int)";
+by (zdiv_undefined_case_tac "a = Numeral0" 1);
by (asm_simp_tac (simpset() addsimps [quorem_div_mod RS self_remainder]) 1);
qed "zmod_self";
Addsimps [zmod_self];
@@ -407,65 +407,65 @@
(*** Computation of division and remainder ***)
-Goal "(#0::int) div b = #0";
+Goal "(Numeral0::int) div b = Numeral0";
by (simp_tac (simpset() addsimps [div_def, divAlg_def]) 1);
qed "zdiv_zero";
-Goal "(#0::int) < b ==> #-1 div b = #-1";
+Goal "(Numeral0::int) < b ==> # -1 div b = # -1";
by (asm_simp_tac (simpset() addsimps [div_def, divAlg_def]) 1);
qed "div_eq_minus1";
-Goal "(#0::int) mod b = #0";
+Goal "(Numeral0::int) mod b = Numeral0";
by (simp_tac (simpset() addsimps [mod_def, divAlg_def]) 1);
qed "zmod_zero";
Addsimps [zdiv_zero, zmod_zero];
-Goal "(#0::int) < b ==> #-1 div b = #-1";
+Goal "(Numeral0::int) < b ==> # -1 div b = # -1";
by (asm_simp_tac (simpset() addsimps [div_def, divAlg_def]) 1);
qed "zdiv_minus1";
-Goal "(#0::int) < b ==> #-1 mod b = b-#1";
+Goal "(Numeral0::int) < b ==> # -1 mod b = b-Numeral1";
by (asm_simp_tac (simpset() addsimps [mod_def, divAlg_def]) 1);
qed "zmod_minus1";
(** a positive, b positive **)
-Goal "[| #0 < a; #0 <= b |] ==> a div b = fst (posDivAlg(a,b))";
+Goal "[| Numeral0 < a; Numeral0 <= b |] ==> a div b = fst (posDivAlg(a,b))";
by (asm_simp_tac (simpset() addsimps [div_def, divAlg_def]) 1);
qed "div_pos_pos";
-Goal "[| #0 < a; #0 <= b |] ==> a mod b = snd (posDivAlg(a,b))";
+Goal "[| Numeral0 < a; Numeral0 <= b |] ==> a mod b = snd (posDivAlg(a,b))";
by (asm_simp_tac (simpset() addsimps [mod_def, divAlg_def]) 1);
qed "mod_pos_pos";
(** a negative, b positive **)
-Goal "[| a < #0; #0 < b |] ==> a div b = fst (negDivAlg(a,b))";
+Goal "[| a < Numeral0; Numeral0 < b |] ==> a div b = fst (negDivAlg(a,b))";
by (asm_simp_tac (simpset() addsimps [div_def, divAlg_def]) 1);
qed "div_neg_pos";
-Goal "[| a < #0; #0 < b |] ==> a mod b = snd (negDivAlg(a,b))";
+Goal "[| a < Numeral0; Numeral0 < b |] ==> a mod b = snd (negDivAlg(a,b))";
by (asm_simp_tac (simpset() addsimps [mod_def, divAlg_def]) 1);
qed "mod_neg_pos";
(** a positive, b negative **)
-Goal "[| #0 < a; b < #0 |] ==> a div b = fst (negateSnd(negDivAlg(-a,-b)))";
+Goal "[| Numeral0 < a; b < Numeral0 |] ==> a div b = fst (negateSnd(negDivAlg(-a,-b)))";
by (asm_simp_tac (simpset() addsimps [div_def, divAlg_def]) 1);
qed "div_pos_neg";
-Goal "[| #0 < a; b < #0 |] ==> a mod b = snd (negateSnd(negDivAlg(-a,-b)))";
+Goal "[| Numeral0 < a; b < Numeral0 |] ==> a mod b = snd (negateSnd(negDivAlg(-a,-b)))";
by (asm_simp_tac (simpset() addsimps [mod_def, divAlg_def]) 1);
qed "mod_pos_neg";
(** a negative, b negative **)
-Goal "[| a < #0; b <= #0 |] ==> a div b = fst (negateSnd(posDivAlg(-a,-b)))";
+Goal "[| a < Numeral0; b <= Numeral0 |] ==> a div b = fst (negateSnd(posDivAlg(-a,-b)))";
by (asm_simp_tac (simpset() addsimps [div_def, divAlg_def]) 1);
qed "div_neg_neg";
-Goal "[| a < #0; b <= #0 |] ==> a mod b = snd (negateSnd(posDivAlg(-a,-b)))";
+Goal "[| a < Numeral0; b <= Numeral0 |] ==> a mod b = snd (negateSnd(posDivAlg(-a,-b)))";
by (asm_simp_tac (simpset() addsimps [mod_def, divAlg_def]) 1);
qed "mod_neg_neg";
@@ -478,28 +478,28 @@
(** Special-case simplification **)
-Goal "a mod (#1::int) = #0";
-by (cut_inst_tac [("a","a"),("b","#1")] pos_mod_sign 1);
-by (cut_inst_tac [("a","a"),("b","#1")] pos_mod_bound 2);
+Goal "a mod (Numeral1::int) = Numeral0";
+by (cut_inst_tac [("a","a"),("b","Numeral1")] pos_mod_sign 1);
+by (cut_inst_tac [("a","a"),("b","Numeral1")] pos_mod_bound 2);
by Auto_tac;
qed "zmod_1";
Addsimps [zmod_1];
-Goal "a div (#1::int) = a";
-by (cut_inst_tac [("a","a"),("b","#1")] zmod_zdiv_equality 1);
+Goal "a div (Numeral1::int) = a";
+by (cut_inst_tac [("a","a"),("b","Numeral1")] zmod_zdiv_equality 1);
by Auto_tac;
qed "zdiv_1";
Addsimps [zdiv_1];
-Goal "a mod (#-1::int) = #0";
-by (cut_inst_tac [("a","a"),("b","#-1")] neg_mod_sign 1);
-by (cut_inst_tac [("a","a"),("b","#-1")] neg_mod_bound 2);
+Goal "a mod (# -1::int) = Numeral0";
+by (cut_inst_tac [("a","a"),("b","# -1")] neg_mod_sign 1);
+by (cut_inst_tac [("a","a"),("b","# -1")] neg_mod_bound 2);
by Auto_tac;
qed "zmod_minus1_right";
Addsimps [zmod_minus1_right];
-Goal "a div (#-1::int) = -a";
-by (cut_inst_tac [("a","a"),("b","#-1")] zmod_zdiv_equality 1);
+Goal "a div (# -1::int) = -a";
+by (cut_inst_tac [("a","a"),("b","# -1")] zmod_zdiv_equality 1);
by Auto_tac;
qed "zdiv_minus1_right";
Addsimps [zdiv_minus1_right];
@@ -507,7 +507,7 @@
(*** Monotonicity in the first argument (divisor) ***)
-Goal "[| a <= a'; #0 < (b::int) |] ==> a div b <= a' div b";
+Goal "[| a <= a'; Numeral0 < (b::int) |] ==> a div b <= a' div b";
by (cut_inst_tac [("a","a"),("b","b")] zmod_zdiv_equality 1);
by (cut_inst_tac [("a","a'"),("b","b")] zmod_zdiv_equality 1);
by (rtac unique_quotient_lemma 1);
@@ -516,7 +516,7 @@
by (ALLGOALS (asm_simp_tac (simpset() addsimps [pos_mod_sign,pos_mod_bound])));
qed "zdiv_mono1";
-Goal "[| a <= a'; (b::int) < #0 |] ==> a' div b <= a div b";
+Goal "[| a <= a'; (b::int) < Numeral0 |] ==> a' div b <= a div b";
by (cut_inst_tac [("a","a"),("b","b")] zmod_zdiv_equality 1);
by (cut_inst_tac [("a","a'"),("b","b")] zmod_zdiv_equality 1);
by (rtac unique_quotient_lemma_neg 1);
@@ -528,14 +528,14 @@
(*** Monotonicity in the second argument (dividend) ***)
-Goal "[| b*q + r = b'*q' + r'; #0 <= b'*q' + r'; \
-\ r' < b'; #0 <= r; #0 < b'; b' <= b |] \
+Goal "[| b*q + r = b'*q' + r'; Numeral0 <= b'*q' + r'; \
+\ r' < b'; Numeral0 <= r; Numeral0 < b'; b' <= b |] \
\ ==> q <= (q'::int)";
-by (subgoal_tac "#0 <= q'" 1);
- by (subgoal_tac "#0 < b'*(q' + #1)" 2);
+by (subgoal_tac "Numeral0 <= q'" 1);
+ by (subgoal_tac "Numeral0 < b'*(q' + Numeral1)" 2);
by (asm_simp_tac (simpset() addsimps [zadd_zmult_distrib2]) 3);
by (asm_full_simp_tac (simpset() addsimps [int_0_less_mult_iff]) 2);
-by (subgoal_tac "b*q < b*(q' + #1)" 1);
+by (subgoal_tac "b*q < b*(q' + Numeral1)" 1);
by (asm_full_simp_tac (simpset() addsimps [zmult_zless_cancel1]) 1);
by (subgoal_tac "b*q = r' - r + b'*q'" 1);
by (Simp_tac 2);
@@ -545,9 +545,9 @@
by Auto_tac;
qed "zdiv_mono2_lemma";
-Goal "[| (#0::int) <= a; #0 < b'; b' <= b |] \
+Goal "[| (Numeral0::int) <= a; Numeral0 < b'; b' <= b |] \
\ ==> a div b <= a div b'";
-by (subgoal_tac "b ~= #0" 1);
+by (subgoal_tac "b ~= Numeral0" 1);
by (arith_tac 2);
by (cut_inst_tac [("a","a"),("b","b")] zmod_zdiv_equality 1);
by (cut_inst_tac [("a","a"),("b","b'")] zmod_zdiv_equality 1);
@@ -557,14 +557,14 @@
by (ALLGOALS (asm_simp_tac (simpset() addsimps [pos_mod_sign,pos_mod_bound])));
qed "zdiv_mono2";
-Goal "[| b*q + r = b'*q' + r'; b'*q' + r' < #0; \
-\ r < b; #0 <= r'; #0 < b'; b' <= b |] \
+Goal "[| b*q + r = b'*q' + r'; b'*q' + r' < Numeral0; \
+\ r < b; Numeral0 <= r'; Numeral0 < b'; b' <= b |] \
\ ==> q' <= (q::int)";
-by (subgoal_tac "q' < #0" 1);
- by (subgoal_tac "b'*q' < #0" 2);
+by (subgoal_tac "q' < Numeral0" 1);
+ by (subgoal_tac "b'*q' < Numeral0" 2);
by (arith_tac 3);
by (asm_full_simp_tac (simpset() addsimps [zmult_less_0_iff]) 2);
-by (subgoal_tac "b*q' < b*(q + #1)" 1);
+by (subgoal_tac "b*q' < b*(q + Numeral1)" 1);
by (asm_full_simp_tac (simpset() addsimps [zmult_zless_cancel1]) 1);
by (asm_simp_tac (simpset() addsimps [zadd_zmult_distrib2]) 1);
by (subgoal_tac "b*q' <= b'*q'" 1);
@@ -574,7 +574,7 @@
by (arith_tac 1);
qed "zdiv_mono2_neg_lemma";
-Goal "[| a < (#0::int); #0 < b'; b' <= b |] \
+Goal "[| a < (Numeral0::int); Numeral0 < b'; b' <= b |] \
\ ==> a div b' <= a div b";
by (cut_inst_tac [("a","a"),("b","b")] zmod_zdiv_equality 1);
by (cut_inst_tac [("a","a"),("b","b'")] zmod_zdiv_equality 1);
@@ -589,7 +589,7 @@
(** proving (a*b) div c = a * (b div c) + a * (b mod c) **)
-Goal "[| quorem((b,c),(q,r)); c ~= #0 |] \
+Goal "[| quorem((b,c),(q,r)); c ~= Numeral0 |] \
\ ==> quorem ((a*b, c), (a*q + a*r div c, a*r mod c))";
by (auto_tac
(claset(),
@@ -602,12 +602,12 @@
val lemma = result();
Goal "(a*b) div c = a*(b div c) + a*(b mod c) div (c::int)";
-by (zdiv_undefined_case_tac "c = #0" 1);
+by (zdiv_undefined_case_tac "c = Numeral0" 1);
by (blast_tac (claset() addIs [quorem_div_mod RS lemma RS quorem_div]) 1);
qed "zdiv_zmult1_eq";
Goal "(a*b) mod c = a*(b mod c) mod (c::int)";
-by (zdiv_undefined_case_tac "c = #0" 1);
+by (zdiv_undefined_case_tac "c = Numeral0" 1);
by (blast_tac (claset() addIs [quorem_div_mod RS lemma RS quorem_mod]) 1);
qed "zmod_zmult1_eq";
@@ -623,27 +623,27 @@
by (rtac zmod_zmult1_eq 1);
qed "zmod_zmult_distrib";
-Goal "b ~= (#0::int) ==> (a*b) div b = a";
+Goal "b ~= (Numeral0::int) ==> (a*b) div b = a";
by (asm_simp_tac (simpset() addsimps [zdiv_zmult1_eq]) 1);
qed "zdiv_zmult_self1";
-Goal "b ~= (#0::int) ==> (b*a) div b = a";
+Goal "b ~= (Numeral0::int) ==> (b*a) div b = a";
by (stac zmult_commute 1 THEN etac zdiv_zmult_self1 1);
qed "zdiv_zmult_self2";
Addsimps [zdiv_zmult_self1, zdiv_zmult_self2];
-Goal "(a*b) mod b = (#0::int)";
+Goal "(a*b) mod b = (Numeral0::int)";
by (simp_tac (simpset() addsimps [zmod_zmult1_eq]) 1);
qed "zmod_zmult_self1";
-Goal "(b*a) mod b = (#0::int)";
+Goal "(b*a) mod b = (Numeral0::int)";
by (simp_tac (simpset() addsimps [zmult_commute, zmod_zmult1_eq]) 1);
qed "zmod_zmult_self2";
Addsimps [zmod_zmult_self1, zmod_zmult_self2];
-Goal "(m mod d = #0) = (EX q::int. m = d*q)";
+Goal "(m mod d = Numeral0) = (EX q::int. m = d*q)";
by (cut_inst_tac [("a","m"),("b","d")] zmod_zdiv_equality 1);
by Auto_tac;
qed "zmod_eq_0_iff";
@@ -652,7 +652,7 @@
(** proving (a+b) div c = a div c + b div c + ((a mod c + b mod c) div c) **)
-Goal "[| quorem((a,c),(aq,ar)); quorem((b,c),(bq,br)); c ~= #0 |] \
+Goal "[| quorem((a,c),(aq,ar)); quorem((b,c),(bq,br)); c ~= Numeral0 |] \
\ ==> quorem ((a+b, c), (aq + bq + (ar+br) div c, (ar+br) mod c))";
by (auto_tac
(claset(),
@@ -666,19 +666,19 @@
(*NOT suitable for rewriting: the RHS has an instance of the LHS*)
Goal "(a+b) div (c::int) = a div c + b div c + ((a mod c + b mod c) div c)";
-by (zdiv_undefined_case_tac "c = #0" 1);
+by (zdiv_undefined_case_tac "c = Numeral0" 1);
by (blast_tac (claset() addIs [[quorem_div_mod,quorem_div_mod]
MRS lemma RS quorem_div]) 1);
qed "zdiv_zadd1_eq";
Goal "(a+b) mod (c::int) = (a mod c + b mod c) mod c";
-by (zdiv_undefined_case_tac "c = #0" 1);
+by (zdiv_undefined_case_tac "c = Numeral0" 1);
by (blast_tac (claset() addIs [[quorem_div_mod,quorem_div_mod]
MRS lemma RS quorem_mod]) 1);
qed "zmod_zadd1_eq";
-Goal "(a mod b) div b = (#0::int)";
-by (zdiv_undefined_case_tac "b = #0" 1);
+Goal "(a mod b) div b = (Numeral0::int)";
+by (zdiv_undefined_case_tac "b = Numeral0" 1);
by (auto_tac (claset(),
simpset() addsimps [linorder_neq_iff,
pos_mod_sign, pos_mod_bound, div_pos_pos_trivial,
@@ -687,7 +687,7 @@
Addsimps [mod_div_trivial];
Goal "(a mod b) mod b = a mod (b::int)";
-by (zdiv_undefined_case_tac "b = #0" 1);
+by (zdiv_undefined_case_tac "b = Numeral0" 1);
by (auto_tac (claset(),
simpset() addsimps [linorder_neq_iff,
pos_mod_sign, pos_mod_bound, mod_pos_pos_trivial,
@@ -710,22 +710,22 @@
qed "zmod_zadd_right_eq";
-Goal "a ~= (#0::int) ==> (a+b) div a = b div a + #1";
+Goal "a ~= (Numeral0::int) ==> (a+b) div a = b div a + Numeral1";
by (asm_simp_tac (simpset() addsimps [zdiv_zadd1_eq]) 1);
qed "zdiv_zadd_self1";
-Goal "a ~= (#0::int) ==> (b+a) div a = b div a + #1";
+Goal "a ~= (Numeral0::int) ==> (b+a) div a = b div a + Numeral1";
by (asm_simp_tac (simpset() addsimps [zdiv_zadd1_eq]) 1);
qed "zdiv_zadd_self2";
Addsimps [zdiv_zadd_self1, zdiv_zadd_self2];
Goal "(a+b) mod a = b mod (a::int)";
-by (zdiv_undefined_case_tac "a = #0" 1);
+by (zdiv_undefined_case_tac "a = Numeral0" 1);
by (asm_simp_tac (simpset() addsimps [zmod_zadd1_eq]) 1);
qed "zmod_zadd_self1";
Goal "(b+a) mod a = b mod (a::int)";
-by (zdiv_undefined_case_tac "a = #0" 1);
+by (zdiv_undefined_case_tac "a = Numeral0" 1);
by (asm_simp_tac (simpset() addsimps [zmod_zadd1_eq]) 1);
qed "zmod_zadd_self2";
Addsimps [zmod_zadd_self1, zmod_zadd_self2];
@@ -739,8 +739,8 @@
(** first, four lemmas to bound the remainder for the cases b<0 and b>0 **)
-Goal "[| (#0::int) < c; b < r; r <= #0 |] ==> b*c < b*(q mod c) + r";
-by (subgoal_tac "b * (c - q mod c) < r * #1" 1);
+Goal "[| (Numeral0::int) < c; b < r; r <= Numeral0 |] ==> b*c < b*(q mod c) + r";
+by (subgoal_tac "b * (c - q mod c) < r * Numeral1" 1);
by (asm_full_simp_tac (simpset() addsimps [zdiff_zmult_distrib2]) 1);
by (rtac order_le_less_trans 1);
by (etac zmult_zless_mono1 2);
@@ -751,20 +751,20 @@
[zadd_commute, add1_zle_eq, pos_mod_bound]));
val lemma1 = result();
-Goal "[| (#0::int) < c; b < r; r <= #0 |] ==> b * (q mod c) + r <= #0";
-by (subgoal_tac "b * (q mod c) <= #0" 1);
+Goal "[| (Numeral0::int) < c; b < r; r <= Numeral0 |] ==> b * (q mod c) + r <= Numeral0";
+by (subgoal_tac "b * (q mod c) <= Numeral0" 1);
by (arith_tac 1);
by (asm_simp_tac (simpset() addsimps [zmult_le_0_iff, pos_mod_sign]) 1);
val lemma2 = result();
-Goal "[| (#0::int) < c; #0 <= r; r < b |] ==> #0 <= b * (q mod c) + r";
-by (subgoal_tac "#0 <= b * (q mod c)" 1);
+Goal "[| (Numeral0::int) < c; Numeral0 <= r; r < b |] ==> Numeral0 <= b * (q mod c) + r";
+by (subgoal_tac "Numeral0 <= b * (q mod c)" 1);
by (arith_tac 1);
by (asm_simp_tac (simpset() addsimps [int_0_le_mult_iff, pos_mod_sign]) 1);
val lemma3 = result();
-Goal "[| (#0::int) < c; #0 <= r; r < b |] ==> b * (q mod c) + r < b * c";
-by (subgoal_tac "r * #1 < b * (c - q mod c)" 1);
+Goal "[| (Numeral0::int) < c; Numeral0 <= r; r < b |] ==> b * (q mod c) + r < b * c";
+by (subgoal_tac "r * Numeral1 < b * (c - q mod c)" 1);
by (asm_full_simp_tac (simpset() addsimps [zdiff_zmult_distrib2]) 1);
by (rtac order_less_le_trans 1);
by (etac zmult_zless_mono1 1);
@@ -775,7 +775,7 @@
[zadd_commute, add1_zle_eq, pos_mod_bound]));
val lemma4 = result();
-Goal "[| quorem ((a,b), (q,r)); b ~= #0; #0 < c |] \
+Goal "[| quorem ((a,b), (q,r)); b ~= Numeral0; Numeral0 < c |] \
\ ==> quorem ((a, b*c), (q div c, b*(q mod c) + r))";
by (auto_tac
(claset(),
@@ -786,15 +786,15 @@
lemma1, lemma2, lemma3, lemma4]));
val lemma = result();
-Goal "(#0::int) < c ==> a div (b*c) = (a div b) div c";
-by (zdiv_undefined_case_tac "b = #0" 1);
+Goal "(Numeral0::int) < c ==> a div (b*c) = (a div b) div c";
+by (zdiv_undefined_case_tac "b = Numeral0" 1);
by (force_tac (claset(),
simpset() addsimps [quorem_div_mod RS lemma RS quorem_div,
zmult_eq_0_iff]) 1);
qed "zdiv_zmult2_eq";
-Goal "(#0::int) < c ==> a mod (b*c) = b*(a div b mod c) + a mod b";
-by (zdiv_undefined_case_tac "b = #0" 1);
+Goal "(Numeral0::int) < c ==> a mod (b*c) = b*(a div b mod c) + a mod b";
+by (zdiv_undefined_case_tac "b = Numeral0" 1);
by (force_tac (claset(),
simpset() addsimps [quorem_div_mod RS lemma RS quorem_mod,
zmult_eq_0_iff]) 1);
@@ -803,26 +803,26 @@
(*** Cancellation of common factors in "div" ***)
-Goal "[| (#0::int) < b; c ~= #0 |] ==> (c*a) div (c*b) = a div b";
+Goal "[| (Numeral0::int) < b; c ~= Numeral0 |] ==> (c*a) div (c*b) = a div b";
by (stac zdiv_zmult2_eq 1);
by Auto_tac;
val lemma1 = result();
-Goal "[| b < (#0::int); c ~= #0 |] ==> (c*a) div (c*b) = a div b";
+Goal "[| b < (Numeral0::int); c ~= Numeral0 |] ==> (c*a) div (c*b) = a div b";
by (subgoal_tac "(c * (-a)) div (c * (-b)) = (-a) div (-b)" 1);
by (rtac lemma1 2);
by Auto_tac;
val lemma2 = result();
-Goal "c ~= (#0::int) ==> (c*a) div (c*b) = a div b";
-by (zdiv_undefined_case_tac "b = #0" 1);
+Goal "c ~= (Numeral0::int) ==> (c*a) div (c*b) = a div b";
+by (zdiv_undefined_case_tac "b = Numeral0" 1);
by (auto_tac
(claset(),
simpset() addsimps [read_instantiate [("x", "b")] linorder_neq_iff,
lemma1, lemma2]));
qed "zdiv_zmult_zmult1";
-Goal "c ~= (#0::int) ==> (a*c) div (b*c) = a div b";
+Goal "c ~= (Numeral0::int) ==> (a*c) div (b*c) = a div b";
by (dtac zdiv_zmult_zmult1 1);
by (auto_tac (claset(), simpset() addsimps [zmult_commute]));
qed "zdiv_zmult_zmult2";
@@ -831,20 +831,20 @@
(*** Distribution of factors over "mod" ***)
-Goal "[| (#0::int) < b; c ~= #0 |] ==> (c*a) mod (c*b) = c * (a mod b)";
+Goal "[| (Numeral0::int) < b; c ~= Numeral0 |] ==> (c*a) mod (c*b) = c * (a mod b)";
by (stac zmod_zmult2_eq 1);
by Auto_tac;
val lemma1 = result();
-Goal "[| b < (#0::int); c ~= #0 |] ==> (c*a) mod (c*b) = c * (a mod b)";
+Goal "[| b < (Numeral0::int); c ~= Numeral0 |] ==> (c*a) mod (c*b) = c * (a mod b)";
by (subgoal_tac "(c * (-a)) mod (c * (-b)) = c * ((-a) mod (-b))" 1);
by (rtac lemma1 2);
by Auto_tac;
val lemma2 = result();
Goal "(c*a) mod (c*b) = (c::int) * (a mod b)";
-by (zdiv_undefined_case_tac "b = #0" 1);
-by (zdiv_undefined_case_tac "c = #0" 1);
+by (zdiv_undefined_case_tac "b = Numeral0" 1);
+by (zdiv_undefined_case_tac "c = Numeral0" 1);
by (auto_tac
(claset(),
simpset() addsimps [read_instantiate [("x", "b")] linorder_neq_iff,
@@ -861,13 +861,13 @@
(** computing "div" by shifting **)
-Goal "(#0::int) <= a ==> (#1 + #2*b) div (#2*a) = b div a";
-by (zdiv_undefined_case_tac "a = #0" 1);
-by (subgoal_tac "#1 <= a" 1);
+Goal "(Numeral0::int) <= a ==> (Numeral1 + # 2*b) div (# 2*a) = b div a";
+by (zdiv_undefined_case_tac "a = Numeral0" 1);
+by (subgoal_tac "Numeral1 <= a" 1);
by (arith_tac 2);
-by (subgoal_tac "#1 < a * #2" 1);
+by (subgoal_tac "Numeral1 < a * # 2" 1);
by (arith_tac 2);
-by (subgoal_tac "#2*(#1 + b mod a) <= #2*a" 1);
+by (subgoal_tac "# 2*(Numeral1 + b mod a) <= # 2*a" 1);
by (rtac zmult_zle_mono2 2);
by (auto_tac (claset(),
simpset() addsimps [zadd_commute, zmult_commute,
@@ -881,18 +881,18 @@
pos_mod_sign RS zadd_zle_mono1 RSN (2,order_trans)]) 1);
by (auto_tac (claset(),
simpset() addsimps [mod_pos_pos_trivial]));
-by (subgoal_tac "#0 <= b mod a" 1);
+by (subgoal_tac "Numeral0 <= b mod a" 1);
by (asm_simp_tac (simpset() addsimps [pos_mod_sign]) 2);
by (arith_tac 1);
qed "pos_zdiv_mult_2";
-Goal "a <= (#0::int) ==> (#1 + #2*b) div (#2*a) = (b+#1) div a";
-by (subgoal_tac "(#1 + #2*(-b-#1)) div (#2 * (-a)) = (-b-#1) div (-a)" 1);
+Goal "a <= (Numeral0::int) ==> (Numeral1 + # 2*b) div (# 2*a) = (b+Numeral1) div a";
+by (subgoal_tac "(Numeral1 + # 2*(-b-Numeral1)) div (# 2 * (-a)) = (-b-Numeral1) div (-a)" 1);
by (rtac pos_zdiv_mult_2 2);
by (auto_tac (claset(),
simpset() addsimps [zmult_zminus_right]));
-by (subgoal_tac "(#-1 - (#2 * b)) = - (#1 + (#2 * b))" 1);
+by (subgoal_tac "(# -1 - (# 2 * b)) = - (Numeral1 + (# 2 * b))" 1);
by (Simp_tac 2);
by (asm_full_simp_tac (HOL_ss
addsimps [zdiv_zminus_zminus, zdiff_def,
@@ -902,17 +902,17 @@
(*Not clear why this must be proved separately; probably number_of causes
simplification problems*)
-Goal "~ #0 <= x ==> x <= (#0::int)";
+Goal "~ Numeral0 <= x ==> x <= (Numeral0::int)";
by Auto_tac;
val lemma = result();
Goal "number_of (v BIT b) div number_of (w BIT False) = \
-\ (if ~b | (#0::int) <= number_of w \
+\ (if ~b | (Numeral0::int) <= number_of w \
\ then number_of v div (number_of w) \
-\ else (number_of v + (#1::int)) div (number_of w))";
+\ else (number_of v + (Numeral1::int)) div (number_of w))";
by (simp_tac (simpset_of Int.thy addsimps [zadd_assoc, number_of_BIT]) 1);
by (asm_simp_tac (simpset()
- delsimps [number_of_reorient]@bin_arith_extra_simps@bin_rel_simps
+ delsimps [thm "number_of_reorient"]@bin_arith_extra_simps@bin_rel_simps
addsimps [zdiv_zmult_zmult1, pos_zdiv_mult_2, lemma,
neg_zdiv_mult_2]) 1);
qed "zdiv_number_of_BIT";
@@ -921,13 +921,13 @@
(** computing "mod" by shifting (proofs resemble those for "div") **)
-Goal "(#0::int) <= a ==> (#1 + #2*b) mod (#2*a) = #1 + #2 * (b mod a)";
-by (zdiv_undefined_case_tac "a = #0" 1);
-by (subgoal_tac "#1 <= a" 1);
+Goal "(Numeral0::int) <= a ==> (Numeral1 + # 2*b) mod (# 2*a) = Numeral1 + # 2 * (b mod a)";
+by (zdiv_undefined_case_tac "a = Numeral0" 1);
+by (subgoal_tac "Numeral1 <= a" 1);
by (arith_tac 2);
-by (subgoal_tac "#1 < a * #2" 1);
+by (subgoal_tac "Numeral1 < a * # 2" 1);
by (arith_tac 2);
-by (subgoal_tac "#2*(#1 + b mod a) <= #2*a" 1);
+by (subgoal_tac "# 2*(Numeral1 + b mod a) <= # 2*a" 1);
by (rtac zmult_zle_mono2 2);
by (auto_tac (claset(),
simpset() addsimps [zadd_commute, zmult_commute,
@@ -941,19 +941,19 @@
pos_mod_sign RS zadd_zle_mono1 RSN (2,order_trans)]) 1);
by (auto_tac (claset(),
simpset() addsimps [mod_pos_pos_trivial]));
-by (subgoal_tac "#0 <= b mod a" 1);
+by (subgoal_tac "Numeral0 <= b mod a" 1);
by (asm_simp_tac (simpset() addsimps [pos_mod_sign]) 2);
by (arith_tac 1);
qed "pos_zmod_mult_2";
-Goal "a <= (#0::int) ==> (#1 + #2*b) mod (#2*a) = #2 * ((b+#1) mod a) - #1";
+Goal "a <= (Numeral0::int) ==> (Numeral1 + # 2*b) mod (# 2*a) = # 2 * ((b+Numeral1) mod a) - Numeral1";
by (subgoal_tac
- "(#1 + #2*(-b-#1)) mod (#2*(-a)) = #1 + #2*((-b-#1) mod (-a))" 1);
+ "(Numeral1 + # 2*(-b-Numeral1)) mod (# 2*(-a)) = Numeral1 + # 2*((-b-Numeral1) mod (-a))" 1);
by (rtac pos_zmod_mult_2 2);
by (auto_tac (claset(),
simpset() addsimps [zmult_zminus_right]));
-by (subgoal_tac "(#-1 - (#2 * b)) = - (#1 + (#2 * b))" 1);
+by (subgoal_tac "(# -1 - (# 2 * b)) = - (Numeral1 + (# 2 * b))" 1);
by (Simp_tac 2);
by (asm_full_simp_tac (HOL_ss
addsimps [zmod_zminus_zminus, zdiff_def,
@@ -964,10 +964,10 @@
Goal "number_of (v BIT b) mod number_of (w BIT False) = \
\ (if b then \
-\ if (#0::int) <= number_of w \
-\ then #2 * (number_of v mod number_of w) + #1 \
-\ else #2 * ((number_of v + (#1::int)) mod number_of w) - #1 \
-\ else #2 * (number_of v mod number_of w))";
+\ if (Numeral0::int) <= number_of w \
+\ then # 2 * (number_of v mod number_of w) + Numeral1 \
+\ else # 2 * ((number_of v + (Numeral1::int)) mod number_of w) - Numeral1 \
+\ else # 2 * (number_of v mod number_of w))";
by (simp_tac (simpset_of Int.thy addsimps [zadd_assoc, number_of_BIT]) 1);
by (asm_simp_tac (simpset()
delsimps bin_arith_extra_simps@bin_rel_simps
@@ -980,20 +980,20 @@
(** Quotients of signs **)
-Goal "[| a < (#0::int); #0 < b |] ==> a div b < #0";
-by (subgoal_tac "a div b <= #-1" 1);
+Goal "[| a < (Numeral0::int); Numeral0 < b |] ==> a div b < Numeral0";
+by (subgoal_tac "a div b <= # -1" 1);
by (Force_tac 1);
by (rtac order_trans 1);
-by (res_inst_tac [("a'","#-1")] zdiv_mono1 1);
+by (res_inst_tac [("a'","# -1")] zdiv_mono1 1);
by (auto_tac (claset(), simpset() addsimps [zdiv_minus1]));
qed "div_neg_pos_less0";
-Goal "[| (#0::int) <= a; b < #0 |] ==> a div b <= #0";
+Goal "[| (Numeral0::int) <= a; b < Numeral0 |] ==> a div b <= Numeral0";
by (dtac zdiv_mono1_neg 1);
by Auto_tac;
qed "div_nonneg_neg_le0";
-Goal "(#0::int) < b ==> (#0 <= a div b) = (#0 <= a)";
+Goal "(Numeral0::int) < b ==> (Numeral0 <= a div b) = (Numeral0 <= a)";
by Auto_tac;
by (dtac zdiv_mono1 2);
by (auto_tac (claset(), simpset() addsimps [linorder_neq_iff]));
@@ -1001,20 +1001,20 @@
by (blast_tac (claset() addIs [div_neg_pos_less0]) 1);
qed "pos_imp_zdiv_nonneg_iff";
-Goal "b < (#0::int) ==> (#0 <= a div b) = (a <= (#0::int))";
+Goal "b < (Numeral0::int) ==> (Numeral0 <= a div b) = (a <= (Numeral0::int))";
by (stac (zdiv_zminus_zminus RS sym) 1);
by (stac pos_imp_zdiv_nonneg_iff 1);
by Auto_tac;
qed "neg_imp_zdiv_nonneg_iff";
(*But not (a div b <= 0 iff a<=0); consider a=1, b=2 when a div b = 0.*)
-Goal "(#0::int) < b ==> (a div b < #0) = (a < #0)";
+Goal "(Numeral0::int) < b ==> (a div b < Numeral0) = (a < Numeral0)";
by (asm_simp_tac (simpset() addsimps [linorder_not_le RS sym,
pos_imp_zdiv_nonneg_iff]) 1);
qed "pos_imp_zdiv_neg_iff";
(*Again the law fails for <=: consider a = -1, b = -2 when a div b = 0*)
-Goal "b < (#0::int) ==> (a div b < #0) = (#0 < a)";
+Goal "b < (Numeral0::int) ==> (a div b < Numeral0) = (Numeral0 < a)";
by (asm_simp_tac (simpset() addsimps [linorder_not_le RS sym,
neg_imp_zdiv_nonneg_iff]) 1);
qed "neg_imp_zdiv_neg_iff";
--- a/src/HOL/Integ/IntDiv.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Integ/IntDiv.thy Fri Oct 05 21:52:39 2001 +0200
@@ -12,27 +12,27 @@
quorem :: "(int*int) * (int*int) => bool"
"quorem == %((a,b), (q,r)).
a = b*q + r &
- (if #0<b then #0<=r & r<b else b<r & r <= #0)"
+ (if Numeral0 < b then Numeral0<=r & r<b else b<r & r <= Numeral0)"
adjust :: "[int, int, int*int] => int*int"
- "adjust a b == %(q,r). if #0 <= r-b then (#2*q + #1, r-b)
- else (#2*q, r)"
+ "adjust a b == %(q,r). if Numeral0 <= r-b then (# 2*q + Numeral1, r-b)
+ else (# 2*q, r)"
(** the division algorithm **)
(*for the case a>=0, b>0*)
consts posDivAlg :: "int*int => int*int"
-recdef posDivAlg "inv_image less_than (%(a,b). nat(a - b + #1))"
+recdef posDivAlg "inv_image less_than (%(a,b). nat(a - b + Numeral1))"
"posDivAlg (a,b) =
- (if (a<b | b<=#0) then (#0,a)
- else adjust a b (posDivAlg(a, #2*b)))"
+ (if (a<b | b<=Numeral0) then (Numeral0,a)
+ else adjust a b (posDivAlg(a, # 2*b)))"
(*for the case a<0, b>0*)
consts negDivAlg :: "int*int => int*int"
recdef negDivAlg "inv_image less_than (%(a,b). nat(- a - b))"
"negDivAlg (a,b) =
- (if (#0<=a+b | b<=#0) then (#-1,a+b)
- else adjust a b (negDivAlg(a, #2*b)))"
+ (if (Numeral0<=a+b | b<=Numeral0) then (# -1,a+b)
+ else adjust a b (negDivAlg(a, # 2*b)))"
(*for the general case b~=0*)
@@ -44,12 +44,12 @@
including the special case a=0, b<0, because negDivAlg requires a<0*)
divAlg :: "int*int => int*int"
"divAlg ==
- %(a,b). if #0<=a then
- if #0<=b then posDivAlg (a,b)
- else if a=#0 then (#0,#0)
+ %(a,b). if Numeral0<=a then
+ if Numeral0<=b then posDivAlg (a,b)
+ else if a=Numeral0 then (Numeral0,Numeral0)
else negateSnd (negDivAlg (-a,-b))
else
- if #0<b then negDivAlg (a,b)
+ if Numeral0<b then negDivAlg (a,b)
else negateSnd (posDivAlg (-a,-b))"
instance
--- a/src/HOL/Integ/IntPower.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Integ/IntPower.ML Fri Oct 05 21:52:39 2001 +0200
@@ -15,7 +15,7 @@
by (rtac (zmod_zmult_distrib RS sym) 1);
qed "zpower_zmod";
-Goal "#1^y = (#1::int)";
+Goal "Numeral1^y = (Numeral1::int)";
by (induct_tac "y" 1);
by Auto_tac;
qed "zpower_1";
--- a/src/HOL/Integ/IntPower.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Integ/IntPower.thy Fri Oct 05 21:52:39 2001 +0200
@@ -12,7 +12,7 @@
int :: {power}
primrec
- power_0 "p ^ 0 = #1"
+ power_0 "p ^ 0 = Numeral1"
power_Suc "p ^ (Suc n) = (p::int) * (p ^ n)"
end
--- a/src/HOL/Integ/NatSimprocs.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Integ/NatSimprocs.ML Fri Oct 05 21:52:39 2001 +0200
@@ -6,16 +6,16 @@
Simprocs for nat numerals (see also nat_simprocs.ML).
*)
-(** For simplifying Suc m - #n **)
+(** For simplifying Suc m - # n **)
-Goal "#0 < n ==> Suc m - n = m - (n - #1)";
+Goal "Numeral0 < n ==> Suc m - n = m - (n - Numeral1)";
by (asm_simp_tac (simpset() addsplits [nat_diff_split]) 1);
qed "Suc_diff_eq_diff_pred";
(*Now just instantiating n to (number_of v) does the right simplification,
but with some redundant inequality tests.*)
Goal "neg (number_of (bin_pred v)) = (number_of v = (0::nat))";
-by (subgoal_tac "neg (number_of (bin_pred v)) = (number_of v < 1')" 1);
+by (subgoal_tac "neg (number_of (bin_pred v)) = (number_of v < Suc 0)" 1);
by (asm_simp_tac (HOL_ss addsimps [less_Suc_eq_le, le_0_eq]) 1);
by (stac less_number_of_Suc 1);
by (Simp_tac 1);
@@ -78,54 +78,54 @@
Addsimps [nat_rec_number_of, nat_rec_add_eq_if];
-(** For simplifying #m - Suc n **)
+(** For simplifying # m - Suc n **)
-Goal "m - Suc n = (m - #1) - n";
+Goal "m - Suc n = (m - Numeral1) - n";
by (simp_tac (numeral_ss addsplits [nat_diff_split]) 1);
qed "diff_Suc_eq_diff_pred";
(*Obsolete because of natdiff_cancel_numerals
Addsimps [inst "m" "number_of ?v" diff_Suc_eq_diff_pred];
- It LOOPS if #1 is being replaced by 1.
+ It LOOPS if Numeral1 is being replaced by 1.
*)
(** Evens and Odds, for Mutilated Chess Board **)
-(*Case analysis on b<#2*)
-Goal "(n::nat) < #2 ==> n = #0 | n = #1";
+(*Case analysis on b<# 2*)
+Goal "(n::nat) < # 2 ==> n = Numeral0 | n = Numeral1";
by (arith_tac 1);
qed "less_2_cases";
-Goal "Suc(Suc(m)) mod #2 = m mod #2";
-by (subgoal_tac "m mod #2 < #2" 1);
+Goal "Suc(Suc(m)) mod # 2 = m mod # 2";
+by (subgoal_tac "m mod # 2 < # 2" 1);
by (Asm_simp_tac 2);
be (less_2_cases RS disjE) 1;
by (ALLGOALS (asm_simp_tac (simpset() addsimps [mod_Suc])));
qed "mod2_Suc_Suc";
Addsimps [mod2_Suc_Suc];
-Goal "!!m::nat. (0 < m mod #2) = (m mod #2 = #1)";
-by (subgoal_tac "m mod #2 < #2" 1);
+Goal "!!m::nat. (0 < m mod # 2) = (m mod # 2 = Numeral1)";
+by (subgoal_tac "m mod # 2 < # 2" 1);
by (Asm_simp_tac 2);
by (auto_tac (claset(), simpset() delsimps [mod_less_divisor]));
qed "mod2_gr_0";
Addsimps [mod2_gr_0, rename_numerals mod2_gr_0];
-(** Removal of small numerals: #0, #1 and (in additive positions) #2 **)
+(** Removal of small numerals: Numeral0, Numeral1 and (in additive positions) # 2 **)
-Goal "#2 + n = Suc (Suc n)";
+Goal "# 2 + n = Suc (Suc n)";
by (Simp_tac 1);
qed "add_2_eq_Suc";
-Goal "n + #2 = Suc (Suc n)";
+Goal "n + # 2 = Suc (Suc n)";
by (Simp_tac 1);
qed "add_2_eq_Suc'";
Addsimps [numeral_0_eq_0, numeral_1_eq_1, add_2_eq_Suc, add_2_eq_Suc'];
(*Can be used to eliminate long strings of Sucs, but not by default*)
-Goal "Suc (Suc (Suc n)) = #3 + n";
+Goal "Suc (Suc (Suc n)) = # 3 + n";
by (Simp_tac 1);
qed "Suc3_eq_add_3";
@@ -136,21 +136,21 @@
We already have some rules to simplify operands smaller than 3.
**)
-Goal "m div (Suc (Suc (Suc n))) = m div (#3+n)";
+Goal "m div (Suc (Suc (Suc n))) = m div (# 3+n)";
by (simp_tac (simpset() addsimps [Suc3_eq_add_3]) 1);
qed "div_Suc_eq_div_add3";
-Goal "m mod (Suc (Suc (Suc n))) = m mod (#3+n)";
+Goal "m mod (Suc (Suc (Suc n))) = m mod (# 3+n)";
by (simp_tac (simpset() addsimps [Suc3_eq_add_3]) 1);
qed "mod_Suc_eq_mod_add3";
Addsimps [div_Suc_eq_div_add3, mod_Suc_eq_mod_add3];
-Goal "(Suc (Suc (Suc m))) div n = (#3+m) div n";
+Goal "(Suc (Suc (Suc m))) div n = (# 3+m) div n";
by (simp_tac (simpset() addsimps [Suc3_eq_add_3]) 1);
qed "Suc_div_eq_add3_div";
-Goal "(Suc (Suc (Suc m))) mod n = (#3+m) mod n";
+Goal "(Suc (Suc (Suc m))) mod n = (# 3+m) mod n";
by (simp_tac (simpset() addsimps [Suc3_eq_add_3]) 1);
qed "Suc_mod_eq_add3_mod";
--- a/src/HOL/Integ/int_arith1.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Integ/int_arith1.ML Fri Oct 05 21:52:39 2001 +0200
@@ -9,15 +9,15 @@
(** Combining of literal coefficients in sums of products **)
-Goal "(x < y) = (x-y < (#0::int))";
+Goal "(x < y) = (x-y < (Numeral0::int))";
by (simp_tac (simpset() addsimps zcompare_rls) 1);
qed "zless_iff_zdiff_zless_0";
-Goal "(x = y) = (x-y = (#0::int))";
+Goal "(x = y) = (x-y = (Numeral0::int))";
by (simp_tac (simpset() addsimps zcompare_rls) 1);
qed "eq_iff_zdiff_eq_0";
-Goal "(x <= y) = (x-y <= (#0::int))";
+Goal "(x <= y) = (x-y <= (Numeral0::int))";
by (simp_tac (simpset() addsimps zcompare_rls) 1);
qed "zle_iff_zdiff_zle_0";
@@ -97,7 +97,7 @@
val uminus_const = Const ("uminus", HOLogic.intT --> HOLogic.intT);
-(*Thus mk_sum[t] yields t+#0; longer sums don't have a trailing zero*)
+(*Thus mk_sum[t] yields t+Numeral0; longer sums don't have a trailing zero*)
fun mk_sum [] = zero
| mk_sum [t,u] = mk_plus (t, u)
| mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
@@ -157,7 +157,7 @@
handle TERM _ => find_first_coeff (t::past) u terms;
-(*Simplify #1*n and n*#1 to n*)
+(*Simplify Numeral1*n and n*Numeral1 to n*)
val add_0s = [zadd_0, zadd_0_right];
val mult_1s = [zmult_1, zmult_1_right, zmult_minus1, zmult_minus1_right];
@@ -279,7 +279,7 @@
structure CombineNumeralsData =
struct
val add = op + : int*int -> int
- val mk_sum = long_mk_sum (*to work for e.g. #2*x + #3*x *)
+ val mk_sum = long_mk_sum (*to work for e.g. # 2*x + # 3*x *)
val dest_sum = dest_sum
val mk_coeff = mk_coeff
val dest_coeff = dest_coeff 1
@@ -318,35 +318,35 @@
set trace_simp;
fun test s = (Goal s; by (Simp_tac 1));
-test "l + #2 + #2 + #2 + (l + #2) + (oo + #2) = (uu::int)";
+test "l + # 2 + # 2 + # 2 + (l + # 2) + (oo + # 2) = (uu::int)";
-test "#2*u = (u::int)";
-test "(i + j + #12 + (k::int)) - #15 = y";
-test "(i + j + #12 + (k::int)) - #5 = y";
+test "# 2*u = (u::int)";
+test "(i + j + # 12 + (k::int)) - # 15 = y";
+test "(i + j + # 12 + (k::int)) - # 5 = y";
test "y - b < (b::int)";
-test "y - (#3*b + c) < (b::int) - #2*c";
+test "y - (# 3*b + c) < (b::int) - # 2*c";
-test "(#2*x - (u*v) + y) - v*#3*u = (w::int)";
-test "(#2*x*u*v + (u*v)*#4 + y) - v*u*#4 = (w::int)";
-test "(#2*x*u*v + (u*v)*#4 + y) - v*u = (w::int)";
-test "u*v - (x*u*v + (u*v)*#4 + y) = (w::int)";
+test "(# 2*x - (u*v) + y) - v*# 3*u = (w::int)";
+test "(# 2*x*u*v + (u*v)*# 4 + y) - v*u*# 4 = (w::int)";
+test "(# 2*x*u*v + (u*v)*# 4 + y) - v*u = (w::int)";
+test "u*v - (x*u*v + (u*v)*# 4 + y) = (w::int)";
-test "(i + j + #12 + (k::int)) = u + #15 + y";
-test "(i + j*#2 + #12 + (k::int)) = j + #5 + y";
+test "(i + j + # 12 + (k::int)) = u + # 15 + y";
+test "(i + j*# 2 + # 12 + (k::int)) = j + # 5 + y";
-test "#2*y + #3*z + #6*w + #2*y + #3*z + #2*u = #2*y' + #3*z' + #6*w' + #2*y' + #3*z' + u + (vv::int)";
+test "# 2*y + # 3*z + # 6*w + # 2*y + # 3*z + # 2*u = # 2*y' + # 3*z' + # 6*w' + # 2*y' + # 3*z' + u + (vv::int)";
test "a + -(b+c) + b = (d::int)";
test "a + -(b+c) - b = (d::int)";
(*negative numerals*)
-test "(i + j + #-2 + (k::int)) - (u + #5 + y) = zz";
-test "(i + j + #-3 + (k::int)) < u + #5 + y";
-test "(i + j + #3 + (k::int)) < u + #-6 + y";
-test "(i + j + #-12 + (k::int)) - #15 = y";
-test "(i + j + #12 + (k::int)) - #-15 = y";
-test "(i + j + #-12 + (k::int)) - #-15 = y";
+test "(i + j + # -2 + (k::int)) - (u + # 5 + y) = zz";
+test "(i + j + # -3 + (k::int)) < u + # 5 + y";
+test "(i + j + # 3 + (k::int)) < u + # -6 + y";
+test "(i + j + # -12 + (k::int)) - # 15 = y";
+test "(i + j + # 12 + (k::int)) - # -15 = y";
+test "(i + j + # -12 + (k::int)) - # -15 = y";
*)
@@ -410,7 +410,7 @@
zmult_1, zmult_1_right,
zmult_minus1, zmult_minus1_right,
zminus_zadd_distrib, zminus_zminus, zmult_assoc,
- IntDef.Zero_def, int_0, zadd_int RS sym, int_Suc];
+ Zero_int_def, int_0, zadd_int RS sym, int_Suc];
val simprocs = [Int_Times_Assoc.conv, Int_Numeral_Simprocs.combine_numerals]@
Int_Numeral_Simprocs.cancel_numerals;
@@ -455,9 +455,9 @@
(* Some test data
Goal "!!a::int. [| a <= b; c <= d; x+y<z |] ==> a+c <= b+d";
by (fast_arith_tac 1);
-Goal "!!a::int. [| a < b; c < d |] ==> a-d+ #2 <= b+(-c)";
+Goal "!!a::int. [| a < b; c < d |] ==> a-d+ # 2 <= b+(-c)";
by (fast_arith_tac 1);
-Goal "!!a::int. [| a < b; c < d |] ==> a+c+ #1 < b+d";
+Goal "!!a::int. [| a < b; c < d |] ==> a+c+ Numeral1 < b+d";
by (fast_arith_tac 1);
Goal "!!a::int. [| a <= b; b+b <= c |] ==> a+a <= c";
by (fast_arith_tac 1);
@@ -465,7 +465,7 @@
\ ==> a+a <= j+j";
by (fast_arith_tac 1);
Goal "!!a::int. [| a+b < i+j; a<b; i<j |] \
-\ ==> a+a - - #-1 < j+j - #3";
+\ ==> a+a - - # -1 < j+j - # 3";
by (fast_arith_tac 1);
Goal "!!a::int. a+b+c <= i+j+k & a<=b & b<=c & i<=j & j<=k --> a+a+a <= k+k+k";
by (arith_tac 1);
@@ -482,6 +482,6 @@
\ ==> a+a+a+a+a+a <= l+l+l+l+i+l";
by (fast_arith_tac 1);
Goal "!!a::int. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
-\ ==> #6*a <= #5*l+i";
+\ ==> # 6*a <= # 5*l+i";
by (fast_arith_tac 1);
*)
--- a/src/HOL/Integ/int_arith2.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Integ/int_arith2.ML Fri Oct 05 21:52:39 2001 +0200
@@ -5,17 +5,17 @@
(** Simplification of inequalities involving numerical constants **)
-Goal "(w <= z - (#1::int)) = (w<(z::int))";
+Goal "(w <= z - (Numeral1::int)) = (w<(z::int))";
by (arith_tac 1);
qed "zle_diff1_eq";
Addsimps [zle_diff1_eq];
-Goal "(w < z + #1) = (w<=(z::int))";
+Goal "(w < z + Numeral1) = (w<=(z::int))";
by (arith_tac 1);
qed "zle_add1_eq_le";
Addsimps [zle_add1_eq_le];
-Goal "(z = z + w) = (w = (#0::int))";
+Goal "(z = z + w) = (w = (Numeral0::int))";
by (arith_tac 1);
qed "zadd_left_cancel0";
Addsimps [zadd_left_cancel0];
@@ -23,13 +23,13 @@
(* nat *)
-Goal "#0 <= z ==> int (nat z) = z";
+Goal "Numeral0 <= z ==> int (nat z) = z";
by (asm_full_simp_tac
(simpset() addsimps [neg_eq_less_0, zle_def, not_neg_nat]) 1);
qed "nat_0_le";
-Goal "z <= #0 ==> nat z = 0";
-by (case_tac "z = #0" 1);
+Goal "z <= Numeral0 ==> nat z = 0";
+by (case_tac "z = Numeral0" 1);
by (asm_simp_tac (simpset() addsimps [nat_le_int0]) 1);
by (asm_full_simp_tac
(simpset() addsimps [neg_eq_less_0, neg_nat, linorder_neq_iff]) 1);
@@ -37,19 +37,19 @@
Addsimps [nat_0_le, nat_le_0];
-val [major,minor] = Goal "[| #0 <= z; !!m. z = int m ==> P |] ==> P";
+val [major,minor] = Goal "[| Numeral0 <= z; !!m. z = int m ==> P |] ==> P";
by (rtac (major RS nat_0_le RS sym RS minor) 1);
qed "nonneg_eq_int";
-Goal "(nat w = m) = (if #0 <= w then w = int m else m=0)";
+Goal "(nat w = m) = (if Numeral0 <= w then w = int m else m=0)";
by Auto_tac;
qed "nat_eq_iff";
-Goal "(m = nat w) = (if #0 <= w then w = int m else m=0)";
+Goal "(m = nat w) = (if Numeral0 <= w then w = int m else m=0)";
by Auto_tac;
qed "nat_eq_iff2";
-Goal "#0 <= w ==> (nat w < m) = (w < int m)";
+Goal "Numeral0 <= w ==> (nat w < m) = (w < int m)";
by (rtac iffI 1);
by (asm_full_simp_tac
(simpset() delsimps [zless_int] addsimps [zless_int RS sym]) 2);
@@ -57,7 +57,7 @@
by (Simp_tac 1);
qed "nat_less_iff";
-Goal "(int m = z) = (m = nat z & #0 <= z)";
+Goal "(int m = z) = (m = nat z & Numeral0 <= z)";
by (auto_tac (claset(), simpset() addsimps [nat_eq_iff2]));
qed "int_eq_iff";
@@ -67,26 +67,26 @@
(*Users don't want to see (int 0), int(Suc 0) or w + - z*)
Addsimps [int_0, int_Suc, symmetric zdiff_def];
-Goal "nat #0 = 0";
+Goal "nat Numeral0 = 0";
by (simp_tac (simpset() addsimps [nat_eq_iff]) 1);
qed "nat_0";
-Goal "nat #1 = 1";
+Goal "nat Numeral1 = Suc 0";
by (simp_tac (simpset() addsimps [nat_eq_iff]) 1);
qed "nat_1";
-Goal "nat #2 = 2";
+Goal "nat # 2 = Suc (Suc 0)";
by (simp_tac (simpset() addsimps [nat_eq_iff]) 1);
qed "nat_2";
-Goal "#0 <= w ==> (nat w < nat z) = (w<z)";
+Goal "Numeral0 <= w ==> (nat w < nat z) = (w<z)";
by (case_tac "neg z" 1);
by (auto_tac (claset(), simpset() addsimps [nat_less_iff]));
by (auto_tac (claset() addIs [zless_trans],
simpset() addsimps [neg_eq_less_0, zle_def]));
qed "nat_less_eq_zless";
-Goal "#0 < w | #0 <= z ==> (nat w <= nat z) = (w<=z)";
+Goal "Numeral0 < w | Numeral0 <= z ==> (nat w <= nat z) = (w<=z)";
by (auto_tac (claset(),
simpset() addsimps [linorder_not_less RS sym,
zless_nat_conj]));
@@ -106,11 +106,11 @@
*)
Goalw [zabs_def]
- "P(abs(i::int)) = ((#0 <= i --> P i) & (i < #0 --> P(-i)))";
+ "P(abs(i::int)) = ((Numeral0 <= i --> P i) & (i < Numeral0 --> P(-i)))";
by(Simp_tac 1);
qed "zabs_split";
-Goal "#0 <= abs (z::int)";
+Goal "Numeral0 <= abs (z::int)";
by (simp_tac (simpset() addsimps [zabs_def]) 1);
qed "zero_le_zabs";
AddIffs [zero_le_zabs];
--- a/src/HOL/Integ/int_factor_simprocs.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Integ/int_factor_simprocs.ML Fri Oct 05 21:52:39 2001 +0200
@@ -10,27 +10,27 @@
(** Factor cancellation theorems for "int" **)
-Goal "!!k::int. (k*m <= k*n) = ((#0 < k --> m<=n) & (k < #0 --> n<=m))";
+Goal "!!k::int. (k*m <= k*n) = ((Numeral0 < k --> m<=n) & (k < Numeral0 --> n<=m))";
by (stac zmult_zle_cancel1 1);
by Auto_tac;
qed "int_mult_le_cancel1";
-Goal "!!k::int. (k*m < k*n) = ((#0 < k & m<n) | (k < #0 & n<m))";
+Goal "!!k::int. (k*m < k*n) = ((Numeral0 < k & m<n) | (k < Numeral0 & n<m))";
by (stac zmult_zless_cancel1 1);
by Auto_tac;
qed "int_mult_less_cancel1";
-Goal "!!k::int. (k*m = k*n) = (k = #0 | m=n)";
+Goal "!!k::int. (k*m = k*n) = (k = Numeral0 | m=n)";
by Auto_tac;
qed "int_mult_eq_cancel1";
-Goal "!!k::int. k~=#0 ==> (k*m) div (k*n) = (m div n)";
+Goal "!!k::int. k~=Numeral0 ==> (k*m) div (k*n) = (m div n)";
by (stac zdiv_zmult_zmult1 1);
by Auto_tac;
qed "int_mult_div_cancel1";
(*For ExtractCommonTermFun, cancelling common factors*)
-Goal "(k*m) div (k*n) = (if k = (#0::int) then #0 else m div n)";
+Goal "(k*m) div (k*n) = (if k = (Numeral0::int) then Numeral0 else m div n)";
by (simp_tac (simpset() addsimps [int_mult_div_cancel1]) 1);
qed "int_mult_div_cancel_disj";
@@ -114,33 +114,33 @@
set trace_simp;
fun test s = (Goal s; by (Simp_tac 1));
-test "#9*x = #12 * (y::int)";
-test "(#9*x) div (#12 * (y::int)) = z";
-test "#9*x < #12 * (y::int)";
-test "#9*x <= #12 * (y::int)";
+test "# 9*x = # 12 * (y::int)";
+test "(# 9*x) div (# 12 * (y::int)) = z";
+test "# 9*x < # 12 * (y::int)";
+test "# 9*x <= # 12 * (y::int)";
-test "#-99*x = #132 * (y::int)";
-test "(#-99*x) div (#132 * (y::int)) = z";
-test "#-99*x < #132 * (y::int)";
-test "#-99*x <= #132 * (y::int)";
+test "# -99*x = # 132 * (y::int)";
+test "(# -99*x) div (# 132 * (y::int)) = z";
+test "# -99*x < # 132 * (y::int)";
+test "# -99*x <= # 132 * (y::int)";
-test "#999*x = #-396 * (y::int)";
-test "(#999*x) div (#-396 * (y::int)) = z";
-test "#999*x < #-396 * (y::int)";
-test "#999*x <= #-396 * (y::int)";
+test "# 999*x = # -396 * (y::int)";
+test "(# 999*x) div (# -396 * (y::int)) = z";
+test "# 999*x < # -396 * (y::int)";
+test "# 999*x <= # -396 * (y::int)";
-test "#-99*x = #-81 * (y::int)";
-test "(#-99*x) div (#-81 * (y::int)) = z";
-test "#-99*x <= #-81 * (y::int)";
-test "#-99*x < #-81 * (y::int)";
+test "# -99*x = # -81 * (y::int)";
+test "(# -99*x) div (# -81 * (y::int)) = z";
+test "# -99*x <= # -81 * (y::int)";
+test "# -99*x < # -81 * (y::int)";
-test "#-2 * x = #-1 * (y::int)";
-test "#-2 * x = -(y::int)";
-test "(#-2 * x) div (#-1 * (y::int)) = z";
-test "#-2 * x < -(y::int)";
-test "#-2 * x <= #-1 * (y::int)";
-test "-x < #-23 * (y::int)";
-test "-x <= #-23 * (y::int)";
+test "# -2 * x = # -1 * (y::int)";
+test "# -2 * x = -(y::int)";
+test "(# -2 * x) div (# -1 * (y::int)) = z";
+test "# -2 * x < -(y::int)";
+test "# -2 * x <= # -1 * (y::int)";
+test "-x < # -23 * (y::int)";
+test "-x <= # -23 * (y::int)";
*)
--- a/src/HOL/Integ/nat_bin.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Integ/nat_bin.ML Fri Oct 05 21:52:39 2001 +0200
@@ -17,16 +17,16 @@
(*These rewrites should one day be re-oriented...*)
-Goal "#0 = (0::nat)";
+Goal "Numeral0 = (0::nat)";
by (simp_tac (HOL_basic_ss addsimps [nat_0, nat_number_of_def]) 1);
qed "numeral_0_eq_0";
-Goal "#1 = (1::nat)";
-by (simp_tac (HOL_basic_ss addsimps [nat_1, nat_number_of_def]) 1);
+Goal "Numeral1 = (1::nat)";
+by (simp_tac (HOL_basic_ss addsimps [nat_1, nat_number_of_def, One_nat_def]) 1);
qed "numeral_1_eq_1";
-Goal "#2 = (2::nat)";
-by (simp_tac (HOL_basic_ss addsimps [nat_2, nat_number_of_def]) 1);
+Goal "# 2 = Suc 1";
+by (simp_tac (HOL_basic_ss addsimps [nat_2, nat_number_of_def, One_nat_def]) 1);
qed "numeral_2_eq_2";
bind_thm ("zero_eq_numeral_0", numeral_0_eq_0 RS sym);
@@ -35,7 +35,7 @@
(*"neg" is used in rewrite rules for binary comparisons*)
Goal "int (number_of v :: nat) = \
-\ (if neg (number_of v) then #0 \
+\ (if neg (number_of v) then Numeral0 \
\ else (number_of v :: int))";
by (simp_tac
(simpset_of Int.thy addsimps [neg_nat, nat_number_of_def,
@@ -54,13 +54,13 @@
(** Successor **)
-Goal "(#0::int) <= z ==> Suc (nat z) = nat (#1 + z)";
+Goal "(Numeral0::int) <= z ==> Suc (nat z) = nat (Numeral1 + z)";
by (rtac sym 1);
by (asm_simp_tac (simpset() addsimps [nat_eq_iff]) 1);
qed "Suc_nat_eq_nat_zadd1";
Goal "Suc (number_of v) = \
-\ (if neg (number_of v) then #1 else number_of (bin_succ v))";
+\ (if neg (number_of v) then Numeral1 else number_of (bin_succ v))";
by (simp_tac
(simpset_of Int.thy addsimps [neg_nat, nat_1, not_neg_eq_ge_0,
nat_number_of_def, int_Suc,
@@ -69,21 +69,21 @@
Addsimps [Suc_nat_number_of];
Goal "Suc (number_of v + n) = \
-\ (if neg (number_of v) then #1+n else number_of (bin_succ v) + n)";
+\ (if neg (number_of v) then Numeral1+n else number_of (bin_succ v) + n)";
by (Simp_tac 1);
qed "Suc_nat_number_of_add";
-Goal "Suc #0 = #1";
+Goal "Suc Numeral0 = Numeral1";
by (Simp_tac 1);
qed "Suc_numeral_0_eq_1";
-Goal "Suc #1 = #2";
+Goal "Suc Numeral1 = # 2";
by (Simp_tac 1);
qed "Suc_numeral_1_eq_2";
(** Addition **)
-Goal "[| (#0::int) <= z; #0 <= z' |] ==> nat (z+z') = nat z + nat z'";
+Goal "[| (Numeral0::int) <= z; Numeral0 <= z' |] ==> nat (z+z') = nat z + nat z'";
by (rtac (inj_int RS injD) 1);
by (asm_simp_tac (simpset() addsimps [zadd_int RS sym]) 1);
qed "nat_add_distrib";
@@ -103,7 +103,7 @@
(** Subtraction **)
-Goal "[| (#0::int) <= z'; z' <= z |] ==> nat (z-z') = nat z - nat z'";
+Goal "[| (Numeral0::int) <= z'; z' <= z |] ==> nat (z-z') = nat z - nat z'";
by (rtac (inj_int RS injD) 1);
by (asm_simp_tac (simpset() addsimps [zdiff_int RS sym, nat_le_eq_zle]) 1);
qed "nat_diff_distrib";
@@ -122,7 +122,7 @@
"(number_of v :: nat) - number_of v' = \
\ (if neg (number_of v') then number_of v \
\ else let d = number_of (bin_add v (bin_minus v')) in \
-\ if neg d then #0 else nat d)";
+\ if neg d then Numeral0 else nat d)";
by (simp_tac
(simpset_of Int.thy delcongs [if_weak_cong]
addsimps [not_neg_eq_ge_0, nat_0,
@@ -134,22 +134,22 @@
(** Multiplication **)
-Goal "(#0::int) <= z ==> nat (z*z') = nat z * nat z'";
-by (case_tac "#0 <= z'" 1);
+Goal "(Numeral0::int) <= z ==> nat (z*z') = nat z * nat z'";
+by (case_tac "Numeral0 <= z'" 1);
by (asm_full_simp_tac (simpset() addsimps [zmult_le_0_iff]) 2);
by (rtac (inj_int RS injD) 1);
by (asm_simp_tac (simpset() addsimps [zmult_int RS sym,
int_0_le_mult_iff]) 1);
qed "nat_mult_distrib";
-Goal "z <= (#0::int) ==> nat(z*z') = nat(-z) * nat(-z')";
+Goal "z <= (Numeral0::int) ==> nat(z*z') = nat(-z) * nat(-z')";
by (rtac trans 1);
by (rtac nat_mult_distrib 2);
by Auto_tac;
qed "nat_mult_distrib_neg";
Goal "(number_of v :: nat) * number_of v' = \
-\ (if neg (number_of v) then #0 else number_of (bin_mult v v'))";
+\ (if neg (number_of v) then Numeral0 else number_of (bin_mult v v'))";
by (simp_tac
(simpset_of Int.thy addsimps [neg_nat, not_neg_eq_ge_0, nat_number_of_def,
nat_mult_distrib RS sym, number_of_mult,
@@ -161,15 +161,15 @@
(** Quotient **)
-Goal "(#0::int) <= z ==> nat (z div z') = nat z div nat z'";
-by (case_tac "#0 <= z'" 1);
+Goal "(Numeral0::int) <= z ==> nat (z div z') = nat z div nat z'";
+by (case_tac "Numeral0 <= z'" 1);
by (auto_tac (claset(),
simpset() addsimps [div_nonneg_neg_le0, DIVISION_BY_ZERO_DIV]));
-by (zdiv_undefined_case_tac "z' = #0" 1);
+by (zdiv_undefined_case_tac "z' = Numeral0" 1);
by (simp_tac (simpset() addsimps [numeral_0_eq_0, DIVISION_BY_ZERO_DIV]) 1);
by (auto_tac (claset() addSEs [nonneg_eq_int], simpset()));
by (rename_tac "m m'" 1);
-by (subgoal_tac "#0 <= int m div int m'" 1);
+by (subgoal_tac "Numeral0 <= int m div int m'" 1);
by (asm_full_simp_tac
(simpset() addsimps [numeral_0_eq_0, pos_imp_zdiv_nonneg_iff]) 2);
by (rtac (inj_int RS injD) 1);
@@ -184,7 +184,7 @@
qed "nat_div_distrib";
Goal "(number_of v :: nat) div number_of v' = \
-\ (if neg (number_of v) then #0 \
+\ (if neg (number_of v) then Numeral0 \
\ else nat (number_of v div number_of v'))";
by (simp_tac
(simpset_of Int.thy addsimps [not_neg_eq_ge_0, nat_number_of_def, neg_nat,
@@ -197,12 +197,12 @@
(** Remainder **)
(*Fails if z'<0: the LHS collapses to (nat z) but the RHS doesn't*)
-Goal "[| (#0::int) <= z; #0 <= z' |] ==> nat (z mod z') = nat z mod nat z'";
-by (zdiv_undefined_case_tac "z' = #0" 1);
+Goal "[| (Numeral0::int) <= z; Numeral0 <= z' |] ==> nat (z mod z') = nat z mod nat z'";
+by (zdiv_undefined_case_tac "z' = Numeral0" 1);
by (simp_tac (simpset() addsimps [numeral_0_eq_0, DIVISION_BY_ZERO_MOD]) 1);
by (auto_tac (claset() addSEs [nonneg_eq_int], simpset()));
by (rename_tac "m m'" 1);
-by (subgoal_tac "#0 <= int m mod int m'" 1);
+by (subgoal_tac "Numeral0 <= int m mod int m'" 1);
by (asm_full_simp_tac
(simpset() addsimps [nat_less_iff, numeral_0_eq_0, pos_mod_sign]) 2);
by (rtac (inj_int RS injD) 1);
@@ -217,7 +217,7 @@
qed "nat_mod_distrib";
Goal "(number_of v :: nat) mod number_of v' = \
-\ (if neg (number_of v) then #0 \
+\ (if neg (number_of v) then Numeral0 \
\ else if neg (number_of v') then number_of v \
\ else nat (number_of v mod number_of v'))";
by (simp_tac
@@ -233,7 +233,7 @@
(** Equals (=) **)
-Goal "[| (#0::int) <= z; #0 <= z' |] ==> (nat z = nat z') = (z=z')";
+Goal "[| (Numeral0::int) <= z; Numeral0 <= z' |] ==> (nat z = nat z') = (z=z')";
by (auto_tac (claset() addSEs [nonneg_eq_int], simpset()));
qed "eq_nat_nat_iff";
@@ -280,22 +280,22 @@
(*** New versions of existing theorems involving 0, 1, 2 ***)
-(*Maps n to #n for n = 0, 1, 2*)
-val numeral_sym_ss =
- HOL_ss addsimps [numeral_0_eq_0 RS sym,
- numeral_1_eq_1 RS sym,
+(*Maps n to # n for n = 0, 1, 2*)
+val numeral_sym_ss =
+ HOL_ss addsimps [numeral_0_eq_0 RS sym,
+ numeral_1_eq_1 RS sym,
numeral_2_eq_2 RS sym,
Suc_numeral_1_eq_2, Suc_numeral_0_eq_1];
fun rename_numerals th = simplify numeral_sym_ss (Thm.transfer (the_context ()) th);
-(*Maps #n to n for n = 0, 1, 2*)
+(*Maps # n to n for n = 0, 1, 2*)
bind_thms ("numerals", [numeral_0_eq_0, numeral_1_eq_1, numeral_2_eq_2]);
val numeral_ss = simpset() addsimps numerals;
(** Nat **)
-Goal "#0 < n ==> n = Suc(n - #1)";
+Goal "Numeral0 < n ==> n = Suc(n - Numeral1)";
by (asm_full_simp_tac numeral_ss 1);
qed "Suc_pred'";
@@ -329,28 +329,28 @@
AddIffs (map rename_numerals [add_is_0, add_gr_0]);
-Goal "Suc n = n + #1";
+Goal "Suc n = n + Numeral1";
by (asm_simp_tac numeral_ss 1);
qed "Suc_eq_add_numeral_1";
(* These two can be useful when m = number_of... *)
-Goal "(m::nat) + n = (if m=#0 then n else Suc ((m - #1) + n))";
+Goal "(m::nat) + n = (if m=Numeral0 then n else Suc ((m - Numeral1) + n))";
by (case_tac "m" 1);
by (ALLGOALS (asm_simp_tac numeral_ss));
qed "add_eq_if";
-Goal "(m::nat) * n = (if m=#0 then #0 else n + ((m - #1) * n))";
+Goal "(m::nat) * n = (if m=Numeral0 then Numeral0 else n + ((m - Numeral1) * n))";
by (case_tac "m" 1);
by (ALLGOALS (asm_simp_tac numeral_ss));
qed "mult_eq_if";
-Goal "(p ^ m :: nat) = (if m=#0 then #1 else p * (p ^ (m - #1)))";
+Goal "(p ^ m :: nat) = (if m=Numeral0 then Numeral1 else p * (p ^ (m - Numeral1)))";
by (case_tac "m" 1);
by (ALLGOALS (asm_simp_tac numeral_ss));
qed "power_eq_if";
-Goal "[| #0<n; #0<m |] ==> m - n < (m::nat)";
+Goal "[| Numeral0<n; Numeral0<m |] ==> m - n < (m::nat)";
by (asm_full_simp_tac (numeral_ss addsimps [diff_less]) 1);
qed "diff_less'";
@@ -375,20 +375,20 @@
(** Power **)
-Goal "(p::nat) ^ #0 = #1";
+Goal "(p::nat) ^ Numeral0 = Numeral1";
by (simp_tac numeral_ss 1);
qed "power_zero";
-Goal "(p::nat) ^ #1 = p";
+Goal "(p::nat) ^ Numeral1 = p";
by (simp_tac numeral_ss 1);
qed "power_one";
Addsimps [power_zero, power_one];
-Goal "(p::nat) ^ #2 = p*p";
+Goal "(p::nat) ^ # 2 = p*p";
by (simp_tac numeral_ss 1);
qed "power_two";
-Goal "#0 < (i::nat) ==> #0 < i^n";
+Goal "Numeral0 < (i::nat) ==> Numeral0 < i^n";
by (asm_simp_tac numeral_ss 1);
qed "zero_less_power'";
Addsimps [zero_less_power'];
@@ -495,9 +495,9 @@
by Auto_tac;
val lemma1 = result();
-Goal "m+m ~= int 1' + n + n";
+Goal "m+m ~= int (Suc 0) + n + n";
by Auto_tac;
-by (dres_inst_tac [("f", "%x. x mod #2")] arg_cong 1);
+by (dres_inst_tac [("f", "%x. x mod # 2")] arg_cong 1);
by (full_simp_tac (simpset() addsimps [zmod_zadd1_eq]) 1);
val lemma2 = result();
@@ -514,7 +514,7 @@
by (res_inst_tac [("x", "number_of v")] spec 1);
by Safe_tac;
by (ALLGOALS Full_simp_tac);
-by (dres_inst_tac [("f", "%x. x mod #2")] arg_cong 1);
+by (dres_inst_tac [("f", "%x. x mod # 2")] arg_cong 1);
by (full_simp_tac (simpset() addsimps [zmod_zadd1_eq]) 1);
qed "eq_number_of_BIT_Pls";
@@ -524,7 +524,7 @@
[number_of_BIT, number_of_Min, eq_commute]) 1);
by (res_inst_tac [("x", "number_of v")] spec 1);
by Auto_tac;
-by (dres_inst_tac [("f", "%x. x mod #2")] arg_cong 1);
+by (dres_inst_tac [("f", "%x. x mod # 2")] arg_cong 1);
by Auto_tac;
qed "eq_number_of_BIT_Min";
@@ -536,7 +536,7 @@
(*** Further lemmas about "nat" ***)
Goal "nat (abs (w * z)) = nat (abs w) * nat (abs z)";
-by (case_tac "z=#0 | w=#0" 1);
+by (case_tac "z=Numeral0 | w=Numeral0" 1);
by Auto_tac;
by (simp_tac (simpset() addsimps [zabs_def, nat_mult_distrib RS sym,
nat_mult_distrib_neg RS sym, zmult_less_0_iff]) 1);
--- a/src/HOL/Integ/nat_simprocs.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Integ/nat_simprocs.ML Fri Oct 05 21:52:39 2001 +0200
@@ -66,19 +66,19 @@
(** For cancel_numeral_factors **)
-Goal "(#0::nat) < k ==> (k*m <= k*n) = (m<=n)";
+Goal "(Numeral0::nat) < k ==> (k*m <= k*n) = (m<=n)";
by Auto_tac;
qed "nat_mult_le_cancel1";
-Goal "(#0::nat) < k ==> (k*m < k*n) = (m<n)";
+Goal "(Numeral0::nat) < k ==> (k*m < k*n) = (m<n)";
by Auto_tac;
qed "nat_mult_less_cancel1";
-Goal "(#0::nat) < k ==> (k*m = k*n) = (m=n)";
+Goal "(Numeral0::nat) < k ==> (k*m = k*n) = (m=n)";
by Auto_tac;
qed "nat_mult_eq_cancel1";
-Goal "(#0::nat) < k ==> (k*m) div (k*n) = (m div n)";
+Goal "(Numeral0::nat) < k ==> (k*m) div (k*n) = (m div n)";
by Auto_tac;
qed "nat_mult_div_cancel1";
@@ -125,7 +125,7 @@
val zero = mk_numeral 0;
val mk_plus = HOLogic.mk_binop "op +";
-(*Thus mk_sum[t] yields t+#0; longer sums don't have a trailing zero*)
+(*Thus mk_sum[t] yields t+Numeral0; longer sums don't have a trailing zero*)
fun mk_sum [] = zero
| mk_sum [t,u] = mk_plus (t, u)
| mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
@@ -158,7 +158,7 @@
val bin_simps = [add_nat_number_of, nat_number_of_add_left,
diff_nat_number_of, le_nat_number_of_eq_not_less,
less_nat_number_of, mult_nat_number_of,
- Let_number_of, nat_number_of] @
+ thm "Let_number_of", nat_number_of] @
bin_arith_simps @ bin_rel_simps;
fun prep_simproc (name, pats, proc) = Simplifier.mk_simproc name pats proc;
@@ -204,11 +204,11 @@
handle TERM _ => find_first_coeff (t::past) u terms;
-(*Simplify #1*n and n*#1 to n*)
+(*Simplify Numeral1*n and n*Numeral1 to n*)
val add_0s = map rename_numerals [add_0, add_0_right];
val mult_1s = map rename_numerals [mult_1, mult_1_right];
-(*Final simplification: cancel + and *; replace #0 by 0 and #1 by 1*)
+(*Final simplification: cancel + and *; replace Numeral0 by 0 and Numeral1 by 1*)
val simplify_meta_eq =
Int_Numeral_Simprocs.simplify_meta_eq
[numeral_0_eq_0, numeral_1_eq_1, add_0, add_0_right,
@@ -319,7 +319,7 @@
structure CombineNumeralsData =
struct
val add = op + : int*int -> int
- val mk_sum = long_mk_sum (*to work for e.g. #2*x + #3*x *)
+ val mk_sum = long_mk_sum (*to work for e.g. # 2*x + # 3*x *)
val dest_sum = restricted_dest_Sucs_sum
val mk_coeff = mk_coeff
val dest_coeff = dest_coeff
@@ -504,62 +504,62 @@
fun test s = (Goal s; by (Simp_tac 1));
(*cancel_numerals*)
-test "l +( #2) + (#2) + #2 + (l + #2) + (oo + #2) = (uu::nat)";
-test "(#2*length xs < #2*length xs + j)";
-test "(#2*length xs < length xs * #2 + j)";
-test "#2*u = (u::nat)";
-test "#2*u = Suc (u)";
-test "(i + j + #12 + (k::nat)) - #15 = y";
-test "(i + j + #12 + (k::nat)) - #5 = y";
-test "Suc u - #2 = y";
-test "Suc (Suc (Suc u)) - #2 = y";
-test "(i + j + #2 + (k::nat)) - 1 = y";
-test "(i + j + #1 + (k::nat)) - 2 = y";
+test "l +( # 2) + (# 2) + # 2 + (l + # 2) + (oo + # 2) = (uu::nat)";
+test "(# 2*length xs < # 2*length xs + j)";
+test "(# 2*length xs < length xs * # 2 + j)";
+test "# 2*u = (u::nat)";
+test "# 2*u = Suc (u)";
+test "(i + j + # 12 + (k::nat)) - # 15 = y";
+test "(i + j + # 12 + (k::nat)) - # 5 = y";
+test "Suc u - # 2 = y";
+test "Suc (Suc (Suc u)) - # 2 = y";
+test "(i + j + # 2 + (k::nat)) - 1 = y";
+test "(i + j + Numeral1 + (k::nat)) - 2 = y";
-test "(#2*x + (u*v) + y) - v*#3*u = (w::nat)";
-test "(#2*x*u*v + #5 + (u*v)*#4 + y) - v*u*#4 = (w::nat)";
-test "(#2*x*u*v + (u*v)*#4 + y) - v*u = (w::nat)";
-test "Suc (Suc (#2*x*u*v + u*#4 + y)) - u = w";
-test "Suc ((u*v)*#4) - v*#3*u = w";
-test "Suc (Suc ((u*v)*#3)) - v*#3*u = w";
+test "(# 2*x + (u*v) + y) - v*# 3*u = (w::nat)";
+test "(# 2*x*u*v + # 5 + (u*v)*# 4 + y) - v*u*# 4 = (w::nat)";
+test "(# 2*x*u*v + (u*v)*# 4 + y) - v*u = (w::nat)";
+test "Suc (Suc (# 2*x*u*v + u*# 4 + y)) - u = w";
+test "Suc ((u*v)*# 4) - v*# 3*u = w";
+test "Suc (Suc ((u*v)*# 3)) - v*# 3*u = w";
-test "(i + j + #12 + (k::nat)) = u + #15 + y";
-test "(i + j + #32 + (k::nat)) - (u + #15 + y) = zz";
-test "(i + j + #12 + (k::nat)) = u + #5 + y";
+test "(i + j + # 12 + (k::nat)) = u + # 15 + y";
+test "(i + j + # 32 + (k::nat)) - (u + # 15 + y) = zz";
+test "(i + j + # 12 + (k::nat)) = u + # 5 + y";
(*Suc*)
-test "(i + j + #12 + k) = Suc (u + y)";
-test "Suc (Suc (Suc (Suc (Suc (u + y))))) <= ((i + j) + #41 + k)";
-test "(i + j + #5 + k) < Suc (Suc (Suc (Suc (Suc (u + y)))))";
-test "Suc (Suc (Suc (Suc (Suc (u + y))))) - #5 = v";
-test "(i + j + #5 + k) = Suc (Suc (Suc (Suc (Suc (Suc (Suc (u + y)))))))";
-test "#2*y + #3*z + #2*u = Suc (u)";
-test "#2*y + #3*z + #6*w + #2*y + #3*z + #2*u = Suc (u)";
-test "#2*y + #3*z + #6*w + #2*y + #3*z + #2*u = #2*y' + #3*z' + #6*w' + #2*y' + #3*z' + u + (vv::nat)";
-test "#6 + #2*y + #3*z + #4*u = Suc (vv + #2*u + z)";
-test "(#2*n*m) < (#3*(m*n)) + (u::nat)";
+test "(i + j + # 12 + k) = Suc (u + y)";
+test "Suc (Suc (Suc (Suc (Suc (u + y))))) <= ((i + j) + # 41 + k)";
+test "(i + j + # 5 + k) < Suc (Suc (Suc (Suc (Suc (u + y)))))";
+test "Suc (Suc (Suc (Suc (Suc (u + y))))) - # 5 = v";
+test "(i + j + # 5 + k) = Suc (Suc (Suc (Suc (Suc (Suc (Suc (u + y)))))))";
+test "# 2*y + # 3*z + # 2*u = Suc (u)";
+test "# 2*y + # 3*z + # 6*w + # 2*y + # 3*z + # 2*u = Suc (u)";
+test "# 2*y + # 3*z + # 6*w + # 2*y + # 3*z + # 2*u = # 2*y' + # 3*z' + # 6*w' + # 2*y' + # 3*z' + u + (vv::nat)";
+test "# 6 + # 2*y + # 3*z + # 4*u = Suc (vv + # 2*u + z)";
+test "(# 2*n*m) < (# 3*(m*n)) + (u::nat)";
(*negative numerals: FAIL*)
-test "(i + j + #-23 + (k::nat)) < u + #15 + y";
-test "(i + j + #3 + (k::nat)) < u + #-15 + y";
-test "(i + j + #-12 + (k::nat)) - #15 = y";
-test "(i + j + #12 + (k::nat)) - #-15 = y";
-test "(i + j + #-12 + (k::nat)) - #-15 = y";
+test "(i + j + # -23 + (k::nat)) < u + # 15 + y";
+test "(i + j + # 3 + (k::nat)) < u + # -15 + y";
+test "(i + j + # -12 + (k::nat)) - # 15 = y";
+test "(i + j + # 12 + (k::nat)) - # -15 = y";
+test "(i + j + # -12 + (k::nat)) - # -15 = y";
(*combine_numerals*)
-test "k + #3*k = (u::nat)";
-test "Suc (i + #3) = u";
-test "Suc (i + j + #3 + k) = u";
-test "k + j + #3*k + j = (u::nat)";
-test "Suc (j*i + i + k + #5 + #3*k + i*j*#4) = (u::nat)";
-test "(#2*n*m) + (#3*(m*n)) = (u::nat)";
+test "k + # 3*k = (u::nat)";
+test "Suc (i + # 3) = u";
+test "Suc (i + j + # 3 + k) = u";
+test "k + j + # 3*k + j = (u::nat)";
+test "Suc (j*i + i + k + # 5 + # 3*k + i*j*# 4) = (u::nat)";
+test "(# 2*n*m) + (# 3*(m*n)) = (u::nat)";
(*negative numerals: FAIL*)
-test "Suc (i + j + #-3 + k) = u";
+test "Suc (i + j + # -3 + k) = u";
(*cancel_numeral_factors*)
-test "#9*x = #12 * (y::nat)";
-test "(#9*x) div (#12 * (y::nat)) = z";
-test "#9*x < #12 * (y::nat)";
-test "#9*x <= #12 * (y::nat)";
+test "# 9*x = # 12 * (y::nat)";
+test "(# 9*x) div (# 12 * (y::nat)) = z";
+test "# 9*x < # 12 * (y::nat)";
+test "# 9*x <= # 12 * (y::nat)";
(*cancel_factor*)
test "x*k = k*(y::nat)";
@@ -597,7 +597,7 @@
Suc_eq_number_of,eq_number_of_Suc,
mult_0, mult_0_right, mult_Suc, mult_Suc_right,
eq_number_of_0, eq_0_number_of, less_0_number_of,
- nat_number_of, Let_number_of, if_True, if_False];
+ nat_number_of, thm "Let_number_of", if_True, if_False];
val simprocs = [Nat_Times_Assoc.conv,
Nat_Numeral_Simprocs.combine_numerals]@
--- a/src/HOL/Isar_examples/Fibonacci.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Isar_examples/Fibonacci.thy Fri Oct 05 21:52:39 2001 +0200
@@ -29,7 +29,7 @@
consts fib :: "nat => nat"
recdef fib less_than
"fib 0 = 0"
- "fib 1' = 1"
+ "fib (Suc 0) = 1"
"fib (Suc (Suc x)) = fib x + fib (Suc x)"
lemma [simp]: "0 < fib (Suc n)"
@@ -39,7 +39,7 @@
text {* Alternative induction rule. *}
theorem fib_induct:
- "P 0 ==> P 1 ==> (!!n. P (n + 1) ==> P n ==> P (n + 2)) ==> P n"
+ "P 0 ==> P 1 ==> (!!n. P (n + 1) ==> P n ==> P (n + # 2)) ==> P (n::nat)"
by (induct rule: fib.induct, simp+)
@@ -56,7 +56,7 @@
show "?P 0" by simp
show "?P 1" by simp
fix n
- have "fib (n + 2 + k + 1)
+ have "fib (n + # 2 + k + 1)
= fib (n + k + 1) + fib (n + 1 + k + 1)" by simp
also assume "fib (n + k + 1)
= fib (k + 1) * fib (n + 1) + fib k * fib n"
@@ -65,9 +65,9 @@
= fib (k + 1) * fib (n + 1 + 1) + fib k * fib (n + 1)"
(is " _ = ?R2")
also have "?R1 + ?R2
- = fib (k + 1) * fib (n + 2 + 1) + fib k * fib (n + 2)"
+ = fib (k + 1) * fib (n + # 2 + 1) + fib k * fib (n + # 2)"
by (simp add: add_mult_distrib2)
- finally show "?P (n + 2)" .
+ finally show "?P (n + # 2)" .
qed
lemma gcd_fib_Suc_eq_1: "gcd (fib n, fib (n + 1)) = 1" (is "?P n")
@@ -75,14 +75,14 @@
show "?P 0" by simp
show "?P 1" by simp
fix n
- have "fib (n + 2 + 1) = fib (n + 1) + fib (n + 2)"
+ have "fib (n + # 2 + 1) = fib (n + 1) + fib (n + # 2)"
by simp
- also have "gcd (fib (n + 2), ...) = gcd (fib (n + 2), fib (n + 1))"
+ also have "gcd (fib (n + # 2), ...) = gcd (fib (n + # 2), fib (n + 1))"
by (simp only: gcd_add2')
also have "... = gcd (fib (n + 1), fib (n + 1 + 1))"
by (simp add: gcd_commute)
also assume "... = 1"
- finally show "?P (n + 2)" .
+ finally show "?P (n + # 2)" .
qed
lemma gcd_mult_add: "0 < n ==> gcd (n * k + m, n) = gcd (m, n)"
--- a/src/HOL/Isar_examples/HoareEx.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Isar_examples/HoareEx.thy Fri Oct 05 21:52:39 2001 +0200
@@ -39,7 +39,7 @@
*}
lemma
- "|- .{\<acute>(N_update (2 * \<acute>N)) : .{\<acute>N = #10}.}. \<acute>N := 2 * \<acute>N .{\<acute>N = #10}."
+ "|- .{\<acute>(N_update (# 2 * \<acute>N)) : .{\<acute>N = # 10}.}. \<acute>N := # 2 * \<acute>N .{\<acute>N = # 10}."
by (rule assign)
text {*
@@ -49,13 +49,13 @@
``obvious'' consequences as well.
*}
-lemma "|- .{True}. \<acute>N := #10 .{\<acute>N = #10}."
+lemma "|- .{True}. \<acute>N := # 10 .{\<acute>N = # 10}."
by hoare
-lemma "|- .{2 * \<acute>N = #10}. \<acute>N := 2 * \<acute>N .{\<acute>N = #10}."
+lemma "|- .{# 2 * \<acute>N = # 10}. \<acute>N := # 2 * \<acute>N .{\<acute>N = # 10}."
by hoare
-lemma "|- .{\<acute>N = #5}. \<acute>N := 2 * \<acute>N .{\<acute>N = #10}."
+lemma "|- .{\<acute>N = # 5}. \<acute>N := # 2 * \<acute>N .{\<acute>N = # 10}."
by hoare simp
lemma "|- .{\<acute>N + 1 = a + 1}. \<acute>N := \<acute>N + 1 .{\<acute>N = a + 1}."
@@ -112,7 +112,7 @@
lemma "|- .{\<acute>M = \<acute>N}. \<acute>M := \<acute>M + 1 .{\<acute>M ~= \<acute>N}."
proof -
- have "!!m n. m = n --> m + 1 ~= n"
+ have "!!m n::nat. m = n --> m + 1 ~= n"
-- {* inclusion of assertions expressed in ``pure'' logic, *}
-- {* without mentioning the state space *}
by simp
--- a/src/HOL/Isar_examples/MutilatedCheckerboard.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Isar_examples/MutilatedCheckerboard.thy Fri Oct 05 21:52:39 2001 +0200
@@ -76,7 +76,7 @@
by (simp add: below_def less_Suc_eq) blast
lemma Sigma_Suc2:
- "m = n + 2 ==> A <*> below m =
+ "m = n + # 2 ==> A <*> below m =
(A <*> {n}) Un (A <*> {n + 1}) Un (A <*> below n)"
by (auto simp add: below_def) arith
@@ -87,10 +87,10 @@
constdefs
evnodd :: "(nat * nat) set => nat => (nat * nat) set"
- "evnodd A b == A Int {(i, j). (i + j) mod #2 = b}"
+ "evnodd A b == A Int {(i, j). (i + j) mod # 2 = b}"
lemma evnodd_iff:
- "(i, j): evnodd A b = ((i, j): A & (i + j) mod #2 = b)"
+ "(i, j): evnodd A b = ((i, j): A & (i + j) mod # 2 = b)"
by (simp add: evnodd_def)
lemma evnodd_subset: "evnodd A b <= A"
@@ -112,7 +112,7 @@
by (simp add: evnodd_def)
lemma evnodd_insert: "evnodd (insert (i, j) C) b =
- (if (i + j) mod #2 = b
+ (if (i + j) mod # 2 = b
then insert (i, j) (evnodd C b) else evnodd C b)"
by (simp add: evnodd_def) blast
@@ -128,21 +128,21 @@
vertl: "{(i, j), (i + 1, j)} : domino"
lemma dominoes_tile_row:
- "{i} <*> below (2 * n) : tiling domino"
+ "{i} <*> below (# 2 * n) : tiling domino"
(is "?P n" is "?B n : ?T")
proof (induct n)
show "?P 0" by (simp add: below_0 tiling.empty)
fix n assume hyp: "?P n"
- let ?a = "{i} <*> {2 * n + 1} Un {i} <*> {2 * n}"
+ let ?a = "{i} <*> {# 2 * n + 1} Un {i} <*> {# 2 * n}"
have "?B (Suc n) = ?a Un ?B n"
by (auto simp add: Sigma_Suc Un_assoc)
also have "... : ?T"
proof (rule tiling.Un)
- have "{(i, 2 * n), (i, 2 * n + 1)} : domino"
+ have "{(i, # 2 * n), (i, # 2 * n + 1)} : domino"
by (rule domino.horiz)
- also have "{(i, 2 * n), (i, 2 * n + 1)} = ?a" by blast
+ also have "{(i, # 2 * n), (i, # 2 * n + 1)} = ?a" by blast
finally show "... : domino" .
from hyp show "?B n : ?T" .
show "?a <= - ?B n" by blast
@@ -151,13 +151,13 @@
qed
lemma dominoes_tile_matrix:
- "below m <*> below (2 * n) : tiling domino"
+ "below m <*> below (# 2 * n) : tiling domino"
(is "?P m" is "?B m : ?T")
proof (induct m)
show "?P 0" by (simp add: below_0 tiling.empty)
fix m assume hyp: "?P m"
- let ?t = "{m} <*> below (2 * n)"
+ let ?t = "{m} <*> below (# 2 * n)"
have "?B (Suc m) = ?t Un ?B m" by (simp add: Sigma_Suc)
also have "... : ?T"
@@ -170,9 +170,9 @@
qed
lemma domino_singleton:
- "d : domino ==> b < 2 ==> EX i j. evnodd d b = {(i, j)}"
+ "d : domino ==> b < # 2 ==> EX i j. evnodd d b = {(i, j)}"
proof -
- assume b: "b < 2"
+ assume b: "b < # 2"
assume "d : domino"
thus ?thesis (is "?P d")
proof induct
@@ -227,9 +227,9 @@
and at: "a <= - t"
have card_suc:
- "!!b. b < 2 ==> card (?e (a Un t) b) = Suc (card (?e t b))"
+ "!!b. b < # 2 ==> card (?e (a Un t) b) = Suc (card (?e t b))"
proof -
- fix b assume "b < 2"
+ fix b :: nat assume "b < # 2"
have "?e (a Un t) b = ?e a b Un ?e t b" by (rule evnodd_Un)
also obtain i j where e: "?e a b = {(i, j)}"
proof -
@@ -260,15 +260,15 @@
constdefs
mutilated_board :: "nat => nat => (nat * nat) set"
"mutilated_board m n ==
- below (2 * (m + 1)) <*> below (2 * (n + 1))
- - {(0, 0)} - {(2 * m + 1, 2 * n + 1)}"
+ below (# 2 * (m + 1)) <*> below (# 2 * (n + 1))
+ - {(0, 0)} - {(# 2 * m + 1, # 2 * n + 1)}"
theorem mutil_not_tiling: "mutilated_board m n ~: tiling domino"
proof (unfold mutilated_board_def)
let ?T = "tiling domino"
- let ?t = "below (2 * (m + 1)) <*> below (2 * (n + 1))"
+ let ?t = "below (# 2 * (m + 1)) <*> below (# 2 * (n + 1))"
let ?t' = "?t - {(0, 0)}"
- let ?t'' = "?t' - {(2 * m + 1, 2 * n + 1)}"
+ let ?t'' = "?t' - {(# 2 * m + 1, # 2 * n + 1)}"
show "?t'' ~: ?T"
proof
@@ -282,12 +282,12 @@
note [simp] = evnodd_iff evnodd_empty evnodd_insert evnodd_Diff
have "card (?e ?t'' 0) < card (?e ?t' 0)"
proof -
- have "card (?e ?t' 0 - {(2 * m + 1, 2 * n + 1)})
+ have "card (?e ?t' 0 - {(# 2 * m + 1, # 2 * n + 1)})
< card (?e ?t' 0)"
proof (rule card_Diff1_less)
from _ fin show "finite (?e ?t' 0)"
by (rule finite_subset) auto
- show "(2 * m + 1, 2 * n + 1) : ?e ?t' 0" by simp
+ show "(# 2 * m + 1, # 2 * n + 1) : ?e ?t' 0" by simp
qed
thus ?thesis by simp
qed
--- a/src/HOL/Isar_examples/Summation.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Isar_examples/Summation.thy Fri Oct 05 21:52:39 2001 +0200
@@ -31,14 +31,14 @@
*}
theorem sum_of_naturals:
- "2 * (\<Sum>i < n + 1. i) = n * (n + 1)"
+ "# 2 * (\<Sum>i < n + 1. i) = n * (n + 1)"
(is "?P n" is "?S n = _")
proof (induct n)
show "?P 0" by simp
next
- fix n have "?S (n + 1) = ?S n + 2 * (n + 1)" by simp
+ fix n have "?S (n + 1) = ?S n + # 2 * (n + 1)" by simp
also assume "?S n = n * (n + 1)"
- also have "... + 2 * (n + 1) = (n + 1) * (n + 2)" by simp
+ also have "... + # 2 * (n + 1) = (n + 1) * (n + # 2)" by simp
finally show "?P (Suc n)" by simp
qed
@@ -86,14 +86,14 @@
*}
theorem sum_of_odds:
- "(\<Sum>i < n. 2 * i + 1) = n^2"
+ "(\<Sum>i < n. # 2 * i + 1) = n^Suc (Suc 0)"
(is "?P n" is "?S n = _")
proof (induct n)
show "?P 0" by simp
next
- fix n have "?S (n + 1) = ?S n + 2 * n + 1" by simp
- also assume "?S n = n^2"
- also have "... + 2 * n + 1 = (n + 1)^2" by simp
+ fix n have "?S (n + 1) = ?S n + # 2 * n + 1" by simp
+ also assume "?S n = n^Suc (Suc 0)"
+ also have "... + # 2 * n + 1 = (n + 1)^Suc (Suc 0)" by simp
finally show "?P (Suc n)" by simp
qed
@@ -106,28 +106,28 @@
lemmas distrib = add_mult_distrib add_mult_distrib2
theorem sum_of_squares:
- "#6 * (\<Sum>i < n + 1. i^2) = n * (n + 1) * (2 * n + 1)"
+ "# 6 * (\<Sum>i < n + 1. i^Suc (Suc 0)) = n * (n + 1) * (# 2 * n + 1)"
(is "?P n" is "?S n = _")
proof (induct n)
show "?P 0" by simp
next
- fix n have "?S (n + 1) = ?S n + #6 * (n + 1)^2" by (simp add: distrib)
- also assume "?S n = n * (n + 1) * (2 * n + 1)"
- also have "... + #6 * (n + 1)^2 =
- (n + 1) * (n + 2) * (2 * (n + 1) + 1)" by (simp add: distrib)
+ fix n have "?S (n + 1) = ?S n + # 6 * (n + 1)^Suc (Suc 0)" by (simp add: distrib)
+ also assume "?S n = n * (n + 1) * (# 2 * n + 1)"
+ also have "... + # 6 * (n + 1)^Suc (Suc 0) =
+ (n + 1) * (n + # 2) * (# 2 * (n + 1) + 1)" by (simp add: distrib)
finally show "?P (Suc n)" by simp
qed
theorem sum_of_cubes:
- "#4 * (\<Sum>i < n + 1. i^#3) = (n * (n + 1))^2"
+ "# 4 * (\<Sum>i < n + 1. i^# 3) = (n * (n + 1))^Suc (Suc 0)"
(is "?P n" is "?S n = _")
proof (induct n)
show "?P 0" by (simp add: power_eq_if)
next
- fix n have "?S (n + 1) = ?S n + #4 * (n + 1)^#3"
+ fix n have "?S (n + 1) = ?S n + # 4 * (n + 1)^# 3"
by (simp add: power_eq_if distrib)
- also assume "?S n = (n * (n + 1))^2"
- also have "... + #4 * (n + 1)^#3 = ((n + 1) * ((n + 1) + 1))^2"
+ also assume "?S n = (n * (n + 1))^Suc (Suc 0)"
+ also have "... + # 4 * (n + 1)^# 3 = ((n + 1) * ((n + 1) + 1))^Suc (Suc 0)"
by (simp add: power_eq_if distrib)
finally show "?P (Suc n)" by simp
qed
--- a/src/HOL/Lambda/Type.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Lambda/Type.thy Fri Oct 05 21:52:39 2001 +0200
@@ -59,11 +59,11 @@
subsection {* Some examples *}
-lemma "e |- Abs (Abs (Abs (Var 1 $ (Var 2 $ Var 1 $ Var 0)))) : ?T"
+lemma "e |- Abs (Abs (Abs (Var 1 $ (Var # 2 $ Var 1 $ Var 0)))) : ?T"
apply force
done
-lemma "e |- Abs (Abs (Abs (Var 2 $ Var 0 $ (Var 1 $ Var 0)))) : ?T"
+lemma "e |- Abs (Abs (Abs (Var # 2 $ Var 0 $ (Var 1 $ Var 0)))) : ?T"
apply force
done
@@ -219,7 +219,7 @@
"e |- t : T ==> \<forall>e' i U u.
e = (\<lambda>j. if j < i then e' j
else if j = i then U
- else e' (j-1)) -->
+ else e' (j - 1)) -->
e' |- u : U --> e' |- t[u/i] : T"
apply (erule typing.induct)
apply (intro strip)
--- a/src/HOL/Library/Multiset.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Library/Multiset.thy Fri Oct 05 21:52:39 2001 +0200
@@ -28,7 +28,7 @@
"{#} == Abs_multiset (\<lambda>a. 0)"
single :: "'a => 'a multiset" ("{#_#}")
- "{#a#} == Abs_multiset (\<lambda>b. if b = a then 1' else 0)"
+ "{#a#} == Abs_multiset (\<lambda>b. if b = a then 1 else 0)"
count :: "'a multiset => 'a => nat"
"count == Rep_multiset"
@@ -54,7 +54,7 @@
defs (overloaded)
union_def: "M + N == Abs_multiset (\<lambda>a. Rep_multiset M a + Rep_multiset N a)"
diff_def: "M - N == Abs_multiset (\<lambda>a. Rep_multiset M a - Rep_multiset N a)"
- Zero_def [simp]: "0 == {#}"
+ Zero_multiset_def [simp]: "0 == {#}"
size_def: "size M == setsum (count M) (set_of M)"
@@ -66,7 +66,7 @@
apply (simp add: multiset_def)
done
-lemma only1_in_multiset [simp]: "(\<lambda>b. if b = a then 1' else 0) \<in> multiset"
+lemma only1_in_multiset [simp]: "(\<lambda>b. if b = a then 1 else 0) \<in> multiset"
apply (simp add: multiset_def)
done
@@ -139,7 +139,7 @@
apply (simp add: count_def Mempty_def)
done
-theorem count_single [simp]: "count {#b#} a = (if b = a then 1' else 0)"
+theorem count_single [simp]: "count {#b#} a = (if b = a then 1 else 0)"
apply (simp add: count_def single_def)
done
@@ -319,8 +319,8 @@
subsection {* Induction over multisets *}
lemma setsum_decr:
- "finite F ==> 0 < f a ==>
- setsum (f (a := f a - 1')) F = (if a \<in> F then setsum f F - 1 else setsum f F)"
+ "finite F ==> (0::nat) < f a ==>
+ setsum (f (a := f a - 1)) F = (if a \<in> F then setsum f F - 1 else setsum f F)"
apply (erule finite_induct)
apply auto
apply (drule_tac a = a in mk_disjoint_insert)
@@ -328,7 +328,7 @@
done
lemma rep_multiset_induct_aux:
- "P (\<lambda>a. 0) ==> (!!f b. f \<in> multiset ==> P f ==> P (f (b := f b + 1')))
+ "P (\<lambda>a. (0::nat)) ==> (!!f b. f \<in> multiset ==> P f ==> P (f (b := f b + 1)))
==> \<forall>f. f \<in> multiset --> setsum f {x. 0 < f x} = n --> P f"
proof -
case rule_context
@@ -347,14 +347,14 @@
apply (frule setsum_SucD)
apply clarify
apply (rename_tac a)
- apply (subgoal_tac "finite {x. 0 < (f (a := f a - 1')) x}")
+ apply (subgoal_tac "finite {x. 0 < (f (a := f a - 1)) x}")
prefer 2
apply (rule finite_subset)
prefer 2
apply assumption
apply simp
apply blast
- apply (subgoal_tac "f = (f (a := f a - 1'))(a := (f (a := f a - 1')) a + 1')")
+ apply (subgoal_tac "f = (f (a := f a - 1))(a := (f (a := f a - 1)) a + 1)")
prefer 2
apply (rule ext)
apply (simp (no_asm_simp))
@@ -362,7 +362,7 @@
apply blast
apply (erule allE, erule impE, erule_tac [2] mp)
apply blast
- apply (simp (no_asm_simp) add: setsum_decr del: fun_upd_apply)
+ apply (simp (no_asm_simp) add: setsum_decr del: fun_upd_apply One_nat_def)
apply (subgoal_tac "{x. x \<noteq> a --> 0 < f x} = {x. 0 < f x}")
prefer 2
apply blast
@@ -375,7 +375,7 @@
theorem rep_multiset_induct:
"f \<in> multiset ==> P (\<lambda>a. 0) ==>
- (!!f b. f \<in> multiset ==> P f ==> P (f (b := f b + 1'))) ==> P f"
+ (!!f b. f \<in> multiset ==> P f ==> P (f (b := f b + 1))) ==> P f"
apply (insert rep_multiset_induct_aux)
apply blast
done
@@ -390,7 +390,7 @@
apply (rule Rep_multiset_inverse [THEN subst])
apply (rule Rep_multiset [THEN rep_multiset_induct])
apply (rule prem1)
- apply (subgoal_tac "f (b := f b + 1') = (\<lambda>a. f a + (if a = b then 1' else 0))")
+ apply (subgoal_tac "f (b := f b + 1) = (\<lambda>a. f a + (if a = b then 1 else 0))")
prefer 2
apply (simp add: expand_fun_eq)
apply (erule ssubst)
--- a/src/HOL/Library/Nat_Infinity.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Library/Nat_Infinity.thy Fri Oct 05 21:52:39 2001 +0200
@@ -31,14 +31,14 @@
Infty :: inat ("\<infinity>")
defs
- iZero_def: "0 == Fin 0"
+ Zero_inat_def: "0 == Fin 0"
iSuc_def: "iSuc i == case i of Fin n => Fin (Suc n) | \<infinity> => \<infinity>"
iless_def: "m < n ==
case m of Fin m1 => (case n of Fin n1 => m1 < n1 | \<infinity> => True)
| \<infinity> => False"
ile_def: "(m::inat) \<le> n == \<not> (n < m)"
-lemmas inat_defs = iZero_def iSuc_def iless_def ile_def
+lemmas inat_defs = Zero_inat_def iSuc_def iless_def ile_def
lemmas inat_splits = inat.split inat.split_asm
text {*
--- a/src/HOL/Library/Primes.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Library/Primes.thy Fri Oct 05 21:52:39 2001 +0200
@@ -54,7 +54,7 @@
declare gcd.simps [simp del]
-lemma gcd_1 [simp]: "gcd (m, 1') = 1"
+lemma gcd_1 [simp]: "gcd (m, Suc 0) = 1"
apply (simp add: gcd_non_0)
done
@@ -140,8 +140,8 @@
apply (simp add: gcd_commute [of 0])
done
-lemma gcd_1_left [simp]: "gcd (1', m) = 1"
- apply (simp add: gcd_commute [of "1'"])
+lemma gcd_1_left [simp]: "gcd (Suc 0, m) = 1"
+ apply (simp add: gcd_commute [of "Suc 0"])
done
@@ -194,7 +194,7 @@
apply (blast intro: relprime_dvd_mult prime_imp_relprime)
done
-lemma prime_dvd_square: "p \<in> prime ==> p dvd m^2 ==> p dvd m"
+lemma prime_dvd_square: "p \<in> prime ==> p dvd m^Suc (Suc 0) ==> p dvd m"
apply (auto dest: prime_dvd_mult)
done
--- a/src/HOL/Library/Rational_Numbers.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Library/Rational_Numbers.thy Fri Oct 05 21:52:39 2001 +0200
@@ -17,7 +17,7 @@
typedef fraction = "{(a, b) :: int \<times> int | a b. b \<noteq> 0}"
proof
- show "(0, #1) \<in> ?fraction" by simp
+ show "(0, Numeral1) \<in> ?fraction" by simp
qed
constdefs
@@ -140,7 +140,7 @@
instance fraction :: ord ..
defs (overloaded)
- zero_fraction_def: "0 == fract 0 #1"
+ zero_fraction_def: "0 == fract 0 Numeral1"
add_fraction_def: "Q + R ==
fract (num Q * den R + num R * den Q) (den Q * den R)"
minus_fraction_def: "-Q == fract (-(num Q)) (den Q)"
@@ -386,9 +386,9 @@
le_rat_def: "q \<le> r == fraction_of q \<le> fraction_of r"
less_rat_def: "q < r == q \<le> r \<and> q \<noteq> (r::rat)"
abs_rat_def: "\<bar>q\<bar> == if q < 0 then -q else (q::rat)"
- number_of_rat_def: "number_of b == Fract (number_of b) #1"
+ number_of_rat_def: "number_of b == Fract (number_of b) Numeral1"
-theorem zero_rat: "0 = Fract 0 #1"
+theorem zero_rat: "0 = Fract 0 Numeral1"
by (simp add: zero_rat_def zero_fraction_def rat_of_def Fract_def)
theorem add_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==>
@@ -497,17 +497,17 @@
by (induct q) (simp add: zero_rat minus_rat add_rat eq_rat)
show "q - r = q + (-r)"
by (induct q, induct r) (simp add: add_rat minus_rat diff_rat)
- show "(0::rat) = #0"
+ show "(0::rat) = Numeral0"
by (simp add: zero_rat number_of_rat_def)
show "(q * r) * s = q * (r * s)"
by (induct q, induct r, induct s) (simp add: mult_rat zmult_ac)
show "q * r = r * q"
by (induct q, induct r) (simp add: mult_rat zmult_ac)
- show "#1 * q = q"
+ show "Numeral1 * q = q"
by (induct q) (simp add: number_of_rat_def mult_rat)
show "(q + r) * s = q * s + r * s"
by (induct q, induct r, induct s) (simp add: add_rat mult_rat eq_rat int_distrib)
- show "q \<noteq> 0 ==> inverse q * q = #1"
+ show "q \<noteq> 0 ==> inverse q * q = Numeral1"
by (induct q) (simp add: inverse_rat mult_rat number_of_rat_def zero_rat eq_rat)
show "r \<noteq> 0 ==> q / r = q * inverse r"
by (induct q, induct r) (simp add: mult_rat divide_rat inverse_rat zero_rat eq_rat)
@@ -630,15 +630,15 @@
constdefs
rat :: "int => rat" (* FIXME generalize int to any numeric subtype *)
- "rat z == Fract z #1"
+ "rat z == Fract z Numeral1"
int_set :: "rat set" ("\<int>") (* FIXME generalize rat to any numeric supertype *)
"\<int> == range rat"
lemma rat_inject: "(rat z = rat w) = (z = w)"
proof
assume "rat z = rat w"
- hence "Fract z #1 = Fract w #1" by (unfold rat_def)
- hence "\<lfloor>fract z #1\<rfloor> = \<lfloor>fract w #1\<rfloor>" ..
+ hence "Fract z Numeral1 = Fract w Numeral1" by (unfold rat_def)
+ hence "\<lfloor>fract z Numeral1\<rfloor> = \<lfloor>fract w Numeral1\<rfloor>" ..
thus "z = w" by auto
next
assume "z = w"
--- a/src/HOL/Library/Ring_and_Field.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Library/Ring_and_Field.thy Fri Oct 05 21:52:39 2001 +0200
@@ -18,11 +18,11 @@
left_zero [simp]: "0 + a = a"
left_minus [simp]: "- a + a = 0"
diff_minus: "a - b = a + (-b)"
- zero_number: "0 = #0"
+ zero_number: "0 = Numeral0"
mult_assoc: "(a * b) * c = a * (b * c)"
mult_commute: "a * b = b * a"
- left_one [simp]: "#1 * a = a"
+ left_one [simp]: "Numeral1 * a = a"
left_distrib: "(a + b) * c = a * c + b * c"
@@ -32,7 +32,7 @@
abs_if: "\<bar>a\<bar> = (if a < 0 then -a else a)"
axclass field \<subseteq> ring, inverse
- left_inverse [simp]: "a \<noteq> 0 ==> inverse a * a = #1"
+ left_inverse [simp]: "a \<noteq> 0 ==> inverse a * a = Numeral1"
divide_inverse: "b \<noteq> 0 ==> a / b = a * inverse b"
axclass ordered_field \<subseteq> ordered_ring, field
@@ -86,10 +86,10 @@
subsubsection {* Derived rules for multiplication *}
-lemma right_one [simp]: "a = a * (#1::'a::field)"
+lemma right_one [simp]: "a = a * (Numeral1::'a::field)"
proof -
- have "a = #1 * a" by simp
- also have "... = a * #1" by (simp add: mult_commute)
+ have "a = Numeral1 * a" by simp
+ also have "... = a * Numeral1" by (simp add: mult_commute)
finally show ?thesis .
qed
@@ -102,28 +102,28 @@
theorems ring_mult_ac = mult_assoc mult_commute mult_left_commute
-lemma right_inverse [simp]: "a \<noteq> 0 ==> a * inverse (a::'a::field) = #1"
+lemma right_inverse [simp]: "a \<noteq> 0 ==> a * inverse (a::'a::field) = Numeral1"
proof -
have "a * inverse a = inverse a * a" by (simp add: ring_mult_ac)
also assume "a \<noteq> 0"
- hence "inverse a * a = #1" by simp
+ hence "inverse a * a = Numeral1" by simp
finally show ?thesis .
qed
-lemma right_inverse_eq: "b \<noteq> 0 ==> (a / b = #1) = (a = (b::'a::field))"
+lemma right_inverse_eq: "b \<noteq> 0 ==> (a / b = Numeral1) = (a = (b::'a::field))"
proof
assume neq: "b \<noteq> 0"
{
hence "a = (a / b) * b" by (simp add: divide_inverse ring_mult_ac)
- also assume "a / b = #1"
+ also assume "a / b = Numeral1"
finally show "a = b" by simp
next
assume "a = b"
- with neq show "a / b = #1" by (simp add: divide_inverse)
+ with neq show "a / b = Numeral1" by (simp add: divide_inverse)
}
qed
-lemma divide_self [simp]: "a \<noteq> 0 ==> a / (a::'a::field) = #1"
+lemma divide_self [simp]: "a \<noteq> 0 ==> a / (a::'a::field) = Numeral1"
by (simp add: divide_inverse)
--- a/src/HOL/Library/Ring_and_Field_Example.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Library/Ring_and_Field_Example.thy Fri Oct 05 21:52:39 2001 +0200
@@ -13,8 +13,8 @@
show "i - j = i + (-j)" by simp
show "(i * j) * k = i * (j * k)" by simp
show "i * j = j * i" by simp
- show "#1 * i = i" by simp
- show "0 = (#0::int)" by simp
+ show "Numeral1 * i = i" by simp
+ show "0 = (Numeral0::int)" by simp
show "(i + j) * k = i * k + j * k" by (simp add: int_distrib)
show "i \<le> j ==> k + i \<le> k + j" by simp
show "i < j ==> 0 < k ==> k * i < k * j" by (simp add: zmult_zless_mono2)
--- a/src/HOL/Library/While_Combinator.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Library/While_Combinator.thy Fri Oct 05 21:52:39 2001 +0200
@@ -22,9 +22,9 @@
recdef (permissive) while_aux
"same_fst (\<lambda>b. True) (\<lambda>b. same_fst (\<lambda>c. True) (\<lambda>c.
{(t, s). b s \<and> c s = t \<and>
- \<not> (\<exists>f. f 0 = s \<and> (\<forall>i. b (f i) \<and> c (f i) = f (i + 1)))}))"
+ \<not> (\<exists>f. f (0::nat) = s \<and> (\<forall>i. b (f i) \<and> c (f i) = f (i + 1)))}))"
"while_aux (b, c, s) =
- (if (\<exists>f. f 0 = s \<and> (\<forall>i. b (f i) \<and> c (f i) = f (i + 1)))
+ (if (\<exists>f. f (0::nat) = s \<and> (\<forall>i. b (f i) \<and> c (f i) = f (i + 1)))
then arbitrary
else if b s then while_aux (b, c, c s)
else s)"
@@ -42,11 +42,10 @@
lemma while_aux_unfold:
"while_aux (b, c, s) =
- (if \<exists>f. f 0 = s \<and> (\<forall>i. b (f i) \<and> c (f i) = f (i + 1))
+ (if \<exists>f. f (0::nat) = s \<and> (\<forall>i. b (f i) \<and> c (f i) = f (i + 1))
then arbitrary
else if b s then while_aux (b, c, c s)
else s)"
-thm while_aux.simps
apply (rule while_aux_tc [THEN while_aux.simps [THEN trans]])
apply (rule refl)
done
@@ -136,14 +135,14 @@
theory.}
*}
-theorem "P (lfp (\<lambda>N::int set. {#0} \<union> {(n + #2) mod #6 | n. n \<in> N})) =
- P {#0, #4, #2}"
+theorem "P (lfp (\<lambda>N::int set. {Numeral0} \<union> {(n + # 2) mod # 6 | n. n \<in> N})) =
+ P {Numeral0, # 4, # 2}"
proof -
have aux: "!!f A B. {f n | n. A n \<or> B n} = {f n | n. A n} \<union> {f n | n. B n}"
apply blast
done
show ?thesis
- apply (subst lfp_conv_while [where ?U = "{#0, #1, #2, #3, #4, #5}"])
+ apply (subst lfp_conv_while [where ?U = "{Numeral0, Numeral1, # 2, # 3, # 4, # 5}"])
apply (rule monoI)
apply blast
apply simp
--- a/src/HOL/List.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/List.ML Fri Oct 05 21:52:39 2001 +0200
@@ -1323,7 +1323,7 @@
qed_spec_mp "hd_replicate";
Addsimps [hd_replicate];
-Goal "n ~= 0 --> tl(replicate n x) = replicate (n-1) x";
+Goal "n ~= 0 --> tl(replicate n x) = replicate (n - 1) x";
by (induct_tac "n" 1);
by Auto_tac;
qed_spec_mp "tl_replicate";
@@ -1506,19 +1506,19 @@
AddIffs (map rename_numerals
[length_0_conv, length_greater_0_conv, sum_eq_0_conv]);
-Goal "take n (x#xs) = (if n = #0 then [] else x # take (n-#1) xs)";
+Goal "take n (x#xs) = (if n = Numeral0 then [] else x # take (n - Numeral1) xs)";
by (case_tac "n" 1);
by (ALLGOALS
(asm_simp_tac (simpset() addsimps [numeral_0_eq_0, numeral_1_eq_1])));
qed "take_Cons'";
-Goal "drop n (x#xs) = (if n = #0 then x#xs else drop (n-#1) xs)";
+Goal "drop n (x#xs) = (if n = Numeral0 then x#xs else drop (n - Numeral1) xs)";
by (case_tac "n" 1);
by (ALLGOALS
(asm_simp_tac (simpset() addsimps [numeral_0_eq_0, numeral_1_eq_1])));
qed "drop_Cons'";
-Goal "(x#xs)!n = (if n = #0 then x else xs!(n-#1))";
+Goal "(x#xs)!n = (if n = Numeral0 then x else xs!(n - Numeral1))";
by (case_tac "n" 1);
by (ALLGOALS
(asm_simp_tac (simpset() addsimps [numeral_0_eq_0, numeral_1_eq_1])));
--- a/src/HOL/MicroJava/BV/JVM.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/MicroJava/BV/JVM.thy Fri Oct 05 21:52:39 2001 +0200
@@ -22,7 +22,7 @@
"wt_kil G C pTs rT mxs mxl ins ==
bounded (\<lambda>n. succs (ins!n) n) (size ins) \<and> 0 < size ins \<and>
(let first = Some ([],(OK (Class C))#((map OK pTs))@(replicate mxl Err));
- start = OK first#(replicate (size ins-1) (OK None));
+ start = OK first#(replicate (size ins - 1) (OK None));
result = kiljvm G mxs (1+size pTs+mxl) rT ins start
in \<forall>n < size ins. result!n \<noteq> Err)"
@@ -149,7 +149,7 @@
==> \<exists>phi. wt_method G C pTs rT maxs mxl bs phi"
proof -
let ?start = "OK (Some ([],(OK (Class C))#((map OK pTs))@(replicate mxl Err)))
- #(replicate (size bs-1) (OK None))"
+ #(replicate (size bs - 1) (OK None))"
assume wf: "wf_prog wf_mb G"
assume isclass: "is_class G C"
@@ -318,7 +318,7 @@
by (rule is_bcv_kiljvm)
let ?start = "OK (Some ([],(OK (Class C))#((map OK pTs))@(replicate mxl Err)))
- #(replicate (size bs-1) (OK None))"
+ #(replicate (size bs - 1) (OK None))"
{ fix l x
have "set (replicate l x) \<subseteq> {x}"
--- a/src/HOL/MicroJava/BV/Step.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/MicroJava/BV/Step.thy Fri Oct 05 21:52:39 2001 +0200
@@ -114,26 +114,26 @@
"succs (Invoke C mn fpTs) pc = [pc+1]"
-lemma 1: "2 < length a ==> (\<exists>l l' l'' ls. a = l#l'#l''#ls)"
+lemma 1: "Suc (Suc 0) < length a ==> (\<exists>l l' l'' ls. a = l#l'#l''#ls)"
proof (cases a)
- fix x xs assume "a = x#xs" "2 < length a"
+ fix x xs assume "a = x#xs" "Suc (Suc 0) < length a"
thus ?thesis by - (cases xs, simp, cases "tl xs", auto)
qed auto
-lemma 2: "\<not>(2 < length a) ==> a = [] \<or> (\<exists> l. a = [l]) \<or> (\<exists> l l'. a = [l,l'])"
+lemma 2: "\<not>(Suc (Suc 0) < length a) ==> a = [] \<or> (\<exists> l. a = [l]) \<or> (\<exists> l l'. a = [l,l'])"
proof -;
- assume "\<not>(2 < length a)"
- hence "length a < (Suc 2)" by simp
- hence * : "length a = 0 \<or> length a = 1' \<or> length a = 2"
+ assume "\<not>(Suc (Suc 0) < length a)"
+ hence "length a < Suc (Suc (Suc 0))" by simp
+ hence * : "length a = 0 \<or> length a = Suc 0 \<or> length a = Suc (Suc 0)"
by (auto simp add: less_Suc_eq)
{
fix x
- assume "length x = 1'"
+ assume "length x = Suc 0"
hence "\<exists> l. x = [l]" by - (cases x, auto)
} note 0 = this
- have "length a = 2 ==> \<exists>l l'. a = [l,l']" by (cases a, auto dest: 0)
+ have "length a = Suc (Suc 0) ==> \<exists>l l'. a = [l,l']" by (cases a, auto dest: 0)
with * show ?thesis by (auto dest: 0)
qed
@@ -152,7 +152,7 @@
lemma appStore[simp]:
"(app (Store idx) G maxs rT (Some s)) = (\<exists> ts ST LT. s = (ts#ST,LT) \<and> idx < length LT)"
- by (cases s, cases "2 < length (fst s)", auto dest: 1 2 simp add: app_def)
+ by (cases s, cases "Suc (Suc 0) < length (fst s)", auto dest: 1 2 simp add: app_def)
lemma appLitPush[simp]:
"(app (LitPush v) G maxs rT (Some s)) = (maxs < length (fst s) \<and> typeof (\<lambda>v. None) v \<noteq> None)"
@@ -162,13 +162,13 @@
"(app (Getfield F C) G maxs rT (Some s)) =
(\<exists> oT vT ST LT. s = (oT#ST, LT) \<and> is_class G C \<and>
field (G,C) F = Some (C,vT) \<and> G \<turnstile> oT \<preceq> (Class C))"
- by (cases s, cases "2 < length (fst s)", auto dest!: 1 2 simp add: app_def)
+ by (cases s, cases "Suc (Suc 0) < length (fst s)", auto dest!: 1 2 simp add: app_def)
lemma appPutField[simp]:
"(app (Putfield F C) G maxs rT (Some s)) =
(\<exists> vT vT' oT ST LT. s = (vT#oT#ST, LT) \<and> is_class G C \<and>
field (G,C) F = Some (C, vT') \<and> G \<turnstile> oT \<preceq> (Class C) \<and> G \<turnstile> vT \<preceq> vT')"
- by (cases s, cases "2 < length (fst s)", auto dest!: 1 2 simp add: app_def)
+ by (cases s, cases "Suc (Suc 0) < length (fst s)", auto dest!: 1 2 simp add: app_def)
lemma appNew[simp]:
"(app (New C) G maxs rT (Some s)) = (is_class G C \<and> maxs < length (fst s))"
@@ -181,27 +181,27 @@
lemma appPop[simp]:
"(app Pop G maxs rT (Some s)) = (\<exists>ts ST LT. s = (ts#ST,LT))"
- by (cases s, cases "2 < length (fst s)", auto dest: 1 2 simp add: app_def)
+ by (cases s, cases "Suc (Suc 0) < length (fst s)", auto dest: 1 2 simp add: app_def)
lemma appDup[simp]:
"(app Dup G maxs rT (Some s)) = (\<exists>ts ST LT. s = (ts#ST,LT) \<and> maxs < Suc (length ST))"
- by (cases s, cases "2 < length (fst s)", auto dest: 1 2 simp add: app_def)
+ by (cases s, cases "Suc (Suc 0) < length (fst s)", auto dest: 1 2 simp add: app_def)
lemma appDup_x1[simp]:
"(app Dup_x1 G maxs rT (Some s)) = (\<exists>ts1 ts2 ST LT. s = (ts1#ts2#ST,LT) \<and> maxs < Suc (Suc (length ST)))"
- by (cases s, cases "2 < length (fst s)", auto dest: 1 2 simp add: app_def)
+ by (cases s, cases "Suc (Suc 0) < length (fst s)", auto dest: 1 2 simp add: app_def)
lemma appDup_x2[simp]:
"(app Dup_x2 G maxs rT (Some s)) = (\<exists>ts1 ts2 ts3 ST LT. s = (ts1#ts2#ts3#ST,LT) \<and> maxs < Suc (Suc (Suc (length ST))))"
- by (cases s, cases "2 < length (fst s)", auto dest: 1 2 simp add: app_def)
+ by (cases s, cases "Suc (Suc 0) < length (fst s)", auto dest: 1 2 simp add: app_def)
lemma appSwap[simp]:
"app Swap G maxs rT (Some s) = (\<exists>ts1 ts2 ST LT. s = (ts1#ts2#ST,LT))"
- by (cases s, cases "2 < length (fst s)", auto dest: 1 2 simp add: app_def)
+ by (cases s, cases "Suc (Suc 0) < length (fst s)", auto dest: 1 2 simp add: app_def)
lemma appIAdd[simp]:
@@ -238,12 +238,12 @@
lemma appIfcmpeq[simp]:
"app (Ifcmpeq b) G maxs rT (Some s) = (\<exists>ts1 ts2 ST LT. s = (ts1#ts2#ST,LT) \<and>
((\<exists> p. ts1 = PrimT p \<and> ts2 = PrimT p) \<or> (\<exists>r r'. ts1 = RefT r \<and> ts2 = RefT r')))"
- by (cases s, cases "2 < length (fst s)", auto dest: 1 2 simp add: app_def)
+ by (cases s, cases "Suc (Suc 0) < length (fst s)", auto dest: 1 2 simp add: app_def)
lemma appReturn[simp]:
"app Return G maxs rT (Some s) = (\<exists>T ST LT. s = (T#ST,LT) \<and> (G \<turnstile> T \<preceq> rT))"
- by (cases s, cases "2 < length (fst s)", auto dest: 1 2 simp add: app_def)
+ by (cases s, cases "Suc (Suc 0) < length (fst s)", auto dest: 1 2 simp add: app_def)
lemma appGoto[simp]:
"app (Goto branch) G maxs rT (Some s) = True"
--- a/src/HOL/MicroJava/J/Example.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/MicroJava/J/Example.thy Fri Oct 05 21:52:39 2001 +0200
@@ -100,7 +100,7 @@
[(vee, PrimT Boolean)],
[((foo,[Class Base]),Class Base,foo_Base)]))"
foo_Ext_def:"foo_Ext == ([x],[],Expr( {Ext}Cast Ext
- (LAcc x)..vee:=Lit (Intg #1)),
+ (LAcc x)..vee:=Lit (Intg Numeral1)),
Lit Null)"
ExtC_def: "ExtC == (Ext, (Base ,
[(vee, PrimT Integer)],
@@ -127,7 +127,7 @@
"NP" == "NullPointer"
"tprg" == "[ObjectC, BaseC, ExtC]"
"obj1" <= "(Ext, empty((vee, Base)\<mapsto>Bool False)
- ((vee, Ext )\<mapsto>Intg #0))"
+ ((vee, Ext )\<mapsto>Intg Numeral0))"
"s0" == " Norm (empty, empty)"
"s1" == " Norm (empty(a\<mapsto>obj1),empty(e\<mapsto>Addr a))"
"s2" == " Norm (empty(a\<mapsto>obj1),empty(x\<mapsto>Null)(This\<mapsto>Addr a))"
--- a/src/HOL/MicroJava/J/Value.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/MicroJava/J/Value.thy Fri Oct 05 21:52:39 2001 +0200
@@ -40,7 +40,7 @@
primrec
"defpval Void = Unit"
"defpval Boolean = Bool False"
- "defpval Integer = Intg (#0)"
+ "defpval Integer = Intg (Numeral0)"
primrec
"default_val (PrimT pt) = defpval pt"
--- a/src/HOL/Nat.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Nat.ML Fri Oct 05 21:52:39 2001 +0200
@@ -68,7 +68,7 @@
by Auto_tac;
qed "less_Suc_eq_0_disj";
-val prems = Goal "[| P 0; P(1'); !!k. P k ==> P (Suc (Suc k)) |] ==> P n";
+val prems = Goal "[| P 0; P(Suc 0); !!k. P k ==> P (Suc (Suc k)) |] ==> P n";
by (rtac nat_less_induct 1);
by (case_tac "n" 1);
by (case_tac "nat" 2);
@@ -157,7 +157,7 @@
(* Could be (and is, below) generalized in various ways;
However, none of the generalizations are currently in the simpset,
and I dread to think what happens if I put them in *)
-Goal "0 < n ==> Suc(n-1') = n";
+Goal "0 < n ==> Suc(n - Suc 0) = n";
by (asm_simp_tac (simpset() addsplits [nat.split]) 1);
qed "Suc_pred";
Addsimps [Suc_pred];
@@ -238,12 +238,12 @@
qed "add_is_0";
AddIffs [add_is_0];
-Goal "(m+n=1') = (m=1' & n=0 | m=0 & n=1')";
+Goal "(m+n= Suc 0) = (m= Suc 0 & n=0 | m=0 & n= Suc 0)";
by (case_tac "m" 1);
by (Auto_tac);
qed "add_is_1";
-Goal "(1' = m+n) = (m=1' & n=0 | m=0 & n=1')";
+Goal "(Suc 0 = m+n) = (m=Suc 0 & n=0 | m=0 & n= Suc 0)";
by (case_tac "m" 1);
by (Auto_tac);
qed "one_is_add";
@@ -396,11 +396,11 @@
Addsimps [mult_0_right, mult_Suc_right];
-Goal "1 * n = n";
+Goal "(1::nat) * n = n";
by (Asm_simp_tac 1);
qed "mult_1";
-Goal "n * 1 = n";
+Goal "n * (1::nat) = n";
by (Asm_simp_tac 1);
qed "mult_1_right";
@@ -638,14 +638,14 @@
qed "zero_less_mult_iff";
Addsimps [zero_less_mult_iff];
-Goal "(1' <= m*n) = (1<=m & 1<=n)";
+Goal "(Suc 0 <= m*n) = (1<=m & 1<=n)";
by (induct_tac "m" 1);
by (case_tac "n" 2);
by (ALLGOALS Asm_simp_tac);
qed "one_le_mult_iff";
Addsimps [one_le_mult_iff];
-Goal "(m*n = 1') = (m=1 & n=1)";
+Goal "(m*n = Suc 0) = (m=1 & n=1)";
by (induct_tac "m" 1);
by (Simp_tac 1);
by (induct_tac "n" 1);
@@ -654,7 +654,7 @@
qed "mult_eq_1_iff";
Addsimps [mult_eq_1_iff];
-Goal "(1' = m*n) = (m=1 & n=1)";
+Goal "(Suc 0 = m*n) = (m=1 & n=1)";
by(rtac (mult_eq_1_iff RSN (2,trans)) 1);
by (fast_tac (claset() addss simpset()) 1);
qed "one_eq_mult_iff";
--- a/src/HOL/NatArith.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/NatArith.ML Fri Oct 05 21:52:39 2001 +0200
@@ -96,17 +96,17 @@
(** Lemmas for ex/Factorization **)
-Goal "!!m::nat. [| 1' < n; 1' < m |] ==> 1' < m*n";
+Goal "!!m::nat. [| Suc 0 < n; Suc 0 < m |] ==> Suc 0 < m*n";
by (case_tac "m" 1);
by Auto_tac;
qed "one_less_mult";
-Goal "!!m::nat. [| 1' < n; 1' < m |] ==> n<m*n";
+Goal "!!m::nat. [| Suc 0 < n; Suc 0 < m |] ==> n<m*n";
by (case_tac "m" 1);
by Auto_tac;
qed "n_less_m_mult_n";
-Goal "!!m::nat. [| 1' < n; 1' < m |] ==> n<n*m";
+Goal "!!m::nat. [| Suc 0 < n; Suc 0 < m |] ==> n<n*m";
by (case_tac "m" 1);
by Auto_tac;
qed "n_less_n_mult_m";
--- a/src/HOL/NatDef.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/NatDef.ML Fri Oct 05 21:52:39 2001 +0200
@@ -4,13 +4,13 @@
Copyright 1991 University of Cambridge
*)
-Addsimps [One_def];
+Addsimps [One_nat_def];
val rew = rewrite_rule [symmetric Nat_def];
(*** Induction ***)
-val prems = Goalw [Zero_def,Suc_def]
+val prems = Goalw [Zero_nat_def,Suc_def]
"[| P(0); \
\ !!n. P(n) ==> P(Suc(n)) |] ==> P(n)";
by (rtac (Rep_Nat_inverse RS subst) 1); (*types force good instantiation*)
@@ -53,7 +53,7 @@
(*** Distinctness of constructors ***)
-Goalw [Zero_def,Suc_def] "Suc(m) ~= 0";
+Goalw [Zero_nat_def,Suc_def] "Suc(m) ~= 0";
by (rtac (inj_on_Abs_Nat RS inj_on_contraD) 1);
by (rtac Suc_Rep_not_Zero_Rep 1);
by (REPEAT (resolve_tac [Rep_Nat, rew Nat'.Suc_RepI, rew Nat'.Zero_RepI] 1));
@@ -191,7 +191,7 @@
by (blast_tac (claset() addSEs [less_SucE] addIs [less_trans]) 1);
qed "less_Suc_eq";
-Goal "(n<1) = (n=0)";
+Goal "(n < (1::nat)) = (n = 0)";
by (simp_tac (simpset() addsimps [less_Suc_eq]) 1);
qed "less_one";
AddIffs [less_one];
@@ -462,12 +462,13 @@
qed "zero_reorient";
Addsimps [zero_reorient];
+(*Polymorphic, not just for "nat"*)
Goal "True ==> (1 = x) = (x = 1)";
by Auto_tac;
qed "one_reorient";
Addsimps [one_reorient];
-Goal "True ==> (2 = x) = (x = 2)";
-by Auto_tac;
+Goal "True ==> (Suc (Suc 0) = x) = (x = Suc (Suc 0))"; (* FIXME !? *)
+by Auto_tac;
qed "two_reorient";
Addsimps [two_reorient];
--- a/src/HOL/NatDef.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/NatDef.thy Fri Oct 05 21:52:39 2001 +0200
@@ -47,7 +47,7 @@
nat = "Nat'" (Nat'.Zero_RepI)
instance
- nat :: {ord, zero}
+ nat :: {ord, zero, one}
(* abstract constants and syntax *)
@@ -55,23 +55,13 @@
consts
Suc :: nat => nat
pred_nat :: "(nat * nat) set"
- "1" :: nat ("1")
-
-syntax
- "1'" :: nat ("1'")
- "2" :: nat ("2")
-
-translations
- "1'" == "Suc 0"
- "2" == "Suc 1'"
-
local
defs
- Zero_def "0 == Abs_Nat(Zero_Rep)"
+ Zero_nat_def "0 == Abs_Nat(Zero_Rep)"
Suc_def "Suc == (%n. Abs_Nat(Suc_Rep(Rep_Nat(n))))"
- One_def "1 == 1'"
+ One_nat_def "1 == Suc 0"
(*nat operations*)
pred_nat_def "pred_nat == {(m,n). n = Suc m}"
--- a/src/HOL/NumberTheory/Chinese.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/NumberTheory/Chinese.thy Fri Oct 05 21:52:39 2001 +0200
@@ -45,26 +45,26 @@
defs
m_cond_def:
"m_cond n mf ==
- (\<forall>i. i \<le> n --> #0 < mf i) \<and>
- (\<forall>i j. i \<le> n \<and> j \<le> n \<and> i \<noteq> j --> zgcd (mf i, mf j) = #1)"
+ (\<forall>i. i \<le> n --> Numeral0 < mf i) \<and>
+ (\<forall>i j. i \<le> n \<and> j \<le> n \<and> i \<noteq> j --> zgcd (mf i, mf j) = Numeral1)"
km_cond_def:
- "km_cond n kf mf == \<forall>i. i \<le> n --> zgcd (kf i, mf i) = #1"
+ "km_cond n kf mf == \<forall>i. i \<le> n --> zgcd (kf i, mf i) = Numeral1"
lincong_sol_def:
"lincong_sol n kf bf mf x == \<forall>i. i \<le> n --> zcong (kf i * x) (bf i) (mf i)"
mhf_def:
"mhf mf n i ==
- if i = 0 then funprod mf 1' (n - 1')
- else if i = n then funprod mf 0 (n - 1')
- else funprod mf 0 (i - 1') * funprod mf (Suc i) (n - 1' - i)"
+ if i = 0 then funprod mf (Suc 0) (n - Suc 0)
+ else if i = n then funprod mf 0 (n - Suc 0)
+ else funprod mf 0 (i - Suc 0) * funprod mf (Suc i) (n - Suc 0 - i)"
xilin_sol_def:
"xilin_sol i n kf bf mf ==
if 0 < n \<and> i \<le> n \<and> m_cond n mf \<and> km_cond n kf mf then
- (SOME x. #0 \<le> x \<and> x < mf i \<and> zcong (kf i * mhf mf n i * x) (bf i) (mf i))
- else #0"
+ (SOME x. Numeral0 \<le> x \<and> x < mf i \<and> zcong (kf i * mhf mf n i * x) (bf i) (mf i))
+ else Numeral0"
x_sol_def:
"x_sol n kf bf mf == funsum (\<lambda>i. xilin_sol i n kf bf mf * mhf mf n i) 0 n"
@@ -72,15 +72,15 @@
text {* \medskip @{term funprod} and @{term funsum} *}
-lemma funprod_pos: "(\<forall>i. i \<le> n --> #0 < mf i) ==> #0 < funprod mf 0 n"
+lemma funprod_pos: "(\<forall>i. i \<le> n --> Numeral0 < mf i) ==> Numeral0 < funprod mf 0 n"
apply (induct n)
apply auto
apply (simp add: int_0_less_mult_iff)
done
lemma funprod_zgcd [rule_format (no_asm)]:
- "(\<forall>i. k \<le> i \<and> i \<le> k + l --> zgcd (mf i, mf m) = #1) -->
- zgcd (funprod mf k l, mf m) = #1"
+ "(\<forall>i. k \<le> i \<and> i \<le> k + l --> zgcd (mf i, mf m) = Numeral1) -->
+ zgcd (funprod mf k l, mf m) = Numeral1"
apply (induct l)
apply simp_all
apply (rule impI)+
@@ -110,14 +110,14 @@
done
lemma funsum_zero [rule_format (no_asm)]:
- "(\<forall>i. k \<le> i \<and> i \<le> k + l --> f i = #0) --> (funsum f k l) = #0"
+ "(\<forall>i. k \<le> i \<and> i \<le> k + l --> f i = Numeral0) --> (funsum f k l) = Numeral0"
apply (induct l)
apply auto
done
lemma funsum_oneelem [rule_format (no_asm)]:
"k \<le> j --> j \<le> k + l -->
- (\<forall>i. k \<le> i \<and> i \<le> k + l \<and> i \<noteq> j --> f i = #0) -->
+ (\<forall>i. k \<le> i \<and> i \<le> k + l \<and> i \<noteq> j --> f i = Numeral0) -->
funsum f k l = f j"
apply (induct l)
prefer 2
@@ -127,9 +127,9 @@
apply (subgoal_tac "k = j")
apply (simp_all (no_asm_simp))
apply (case_tac "Suc (k + n) = j")
- apply (subgoal_tac "funsum f k n = #0")
+ apply (subgoal_tac "funsum f k n = Numeral0")
apply (rule_tac [2] funsum_zero)
- apply (subgoal_tac [3] "f (Suc (k + n)) = #0")
+ apply (subgoal_tac [3] "f (Suc (k + n)) = Numeral0")
apply (subgoal_tac [3] "j \<le> k + n")
prefer 4
apply arith
@@ -175,7 +175,7 @@
lemma unique_xi_sol:
"0 < n ==> i \<le> n ==> m_cond n mf ==> km_cond n kf mf
- ==> \<exists>!x. #0 \<le> x \<and> x < mf i \<and> [kf i * mhf mf n i * x = bf i] (mod mf i)"
+ ==> \<exists>!x. Numeral0 \<le> x \<and> x < mf i \<and> [kf i * mhf mf n i * x = bf i] (mod mf i)"
apply (rule zcong_lineq_unique)
apply (tactic {* stac (thm "zgcd_zmult_cancel") 2 *})
apply (unfold m_cond_def km_cond_def mhf_def)
@@ -227,7 +227,7 @@
lemma chinese_remainder:
"0 < n ==> m_cond n mf ==> km_cond n kf mf
- ==> \<exists>!x. #0 \<le> x \<and> x < funprod mf 0 n \<and> lincong_sol n kf bf mf x"
+ ==> \<exists>!x. Numeral0 \<le> x \<and> x < funprod mf 0 n \<and> lincong_sol n kf bf mf x"
apply safe
apply (rule_tac [2] m = "funprod mf 0 n" in zcong_zless_imp_eq)
apply (rule_tac [6] zcong_funprod)
@@ -242,7 +242,7 @@
apply (tactic {* stac (thm "zmod_zmult_distrib" RS sym) 7 *})
apply (tactic {* stac (thm "zcong_zmod" RS sym) 7 *})
apply (subgoal_tac [7]
- "#0 \<le> xilin_sol i n kf bf mf \<and> xilin_sol i n kf bf mf < mf i
+ "Numeral0 \<le> xilin_sol i n kf bf mf \<and> xilin_sol i n kf bf mf < mf i
\<and> [kf i * mhf mf n i * xilin_sol i n kf bf mf = bf i] (mod mf i)")
prefer 7
apply (simp add: zmult_ac)
--- a/src/HOL/NumberTheory/EulerFermat.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/NumberTheory/EulerFermat.thy Fri Oct 05 21:52:39 2001 +0200
@@ -29,33 +29,33 @@
inductive "RsetR m"
intros
empty [simp]: "{} \<in> RsetR m"
- insert: "A \<in> RsetR m ==> zgcd (a, m) = #1 ==>
+ insert: "A \<in> RsetR m ==> zgcd (a, m) = Numeral1 ==>
\<forall>a'. a' \<in> A --> \<not> zcong a a' m ==> insert a A \<in> RsetR m"
recdef BnorRset
"measure ((\<lambda>(a, m). nat a) :: int * int => nat)"
"BnorRset (a, m) =
- (if #0 < a then
- let na = BnorRset (a - #1, m)
- in (if zgcd (a, m) = #1 then insert a na else na)
+ (if Numeral0 < a then
+ let na = BnorRset (a - Numeral1, m)
+ in (if zgcd (a, m) = Numeral1 then insert a na else na)
else {})"
defs
- norRRset_def: "norRRset m == BnorRset (m - #1, m)"
+ norRRset_def: "norRRset m == BnorRset (m - Numeral1, m)"
noXRRset_def: "noXRRset m x == (\<lambda>a. a * x) ` norRRset m"
phi_def: "phi m == card (norRRset m)"
is_RRset_def: "is_RRset A m == A \<in> RsetR m \<and> card A = phi m"
RRset2norRR_def:
"RRset2norRR A m a ==
- (if #1 < m \<and> is_RRset A m \<and> a \<in> A then
+ (if Numeral1 < m \<and> is_RRset A m \<and> a \<in> A then
SOME b. zcong a b m \<and> b \<in> norRRset m
- else #0)"
+ else Numeral0)"
constdefs
zcongm :: "int => int => int => bool"
"zcongm m == \<lambda>a b. zcong a b m"
-lemma abs_eq_1_iff [iff]: "(abs z = (#1::int)) = (z = #1 \<or> z = #-1)"
+lemma abs_eq_1_iff [iff]: "(abs z = (Numeral1::int)) = (z = Numeral1 \<or> z = # -1)"
-- {* LCP: not sure why this lemma is needed now *}
apply (auto simp add: zabs_def)
done
@@ -67,7 +67,7 @@
lemma BnorRset_induct:
"(!!a m. P {} a m) ==>
- (!!a m. #0 < (a::int) ==> P (BnorRset (a - #1, m::int)) (a - #1) m
+ (!!a m. Numeral0 < (a::int) ==> P (BnorRset (a - Numeral1, m::int)) (a - Numeral1) m
==> P (BnorRset(a,m)) a m)
==> P (BnorRset(u,v)) u v"
proof -
@@ -75,7 +75,7 @@
show ?thesis
apply (rule BnorRset.induct)
apply safe
- apply (case_tac [2] "#0 < a")
+ apply (case_tac [2] "Numeral0 < a")
apply (rule_tac [2] rule_context)
apply simp_all
apply (simp_all add: BnorRset.simps rule_context)
@@ -94,7 +94,7 @@
apply (auto dest: Bnor_mem_zle)
done
-lemma Bnor_mem_zg [rule_format]: "b \<in> BnorRset (a, m) --> #0 < b"
+lemma Bnor_mem_zg [rule_format]: "b \<in> BnorRset (a, m) --> Numeral0 < b"
apply (induct a m rule: BnorRset_induct)
prefer 2
apply (subst BnorRset.simps)
@@ -103,7 +103,7 @@
done
lemma Bnor_mem_if [rule_format]:
- "zgcd (b, m) = #1 --> #0 < b --> b \<le> a --> b \<in> BnorRset (a, m)"
+ "zgcd (b, m) = Numeral1 --> Numeral0 < b --> b \<le> a --> b \<in> BnorRset (a, m)"
apply (induct a m rule: BnorRset.induct)
apply auto
apply (case_tac "a = b")
@@ -128,7 +128,7 @@
apply (rule_tac [3] allI)
apply (rule_tac [3] impI)
apply (rule_tac [3] zcong_not)
- apply (subgoal_tac [6] "a' \<le> a - #1")
+ apply (subgoal_tac [6] "a' \<le> a - Numeral1")
apply (rule_tac [7] Bnor_mem_zle)
apply (rule_tac [5] Bnor_mem_zg)
apply auto
@@ -142,13 +142,13 @@
apply auto
done
-lemma aux: "a \<le> b - #1 ==> a < (b::int)"
+lemma aux: "a \<le> b - Numeral1 ==> a < (b::int)"
apply auto
done
lemma norR_mem_unique:
- "#1 < m ==>
- zgcd (a, m) = #1 ==> \<exists>!b. [a = b] (mod m) \<and> b \<in> norRRset m"
+ "Numeral1 < m ==>
+ zgcd (a, m) = Numeral1 ==> \<exists>!b. [a = b] (mod m) \<and> b \<in> norRRset m"
apply (unfold norRRset_def)
apply (cut_tac a = a and m = m in zcong_zless_unique)
apply auto
@@ -158,7 +158,7 @@
apply (rule_tac "x" = "b" in exI)
apply safe
apply (rule Bnor_mem_if)
- apply (case_tac [2] "b = #0")
+ apply (case_tac [2] "b = Numeral0")
apply (auto intro: order_less_le [THEN iffD2])
prefer 2
apply (simp only: zcong_def)
@@ -173,7 +173,7 @@
text {* \medskip @{term noXRRset} *}
lemma RRset_gcd [rule_format]:
- "is_RRset A m ==> a \<in> A --> zgcd (a, m) = #1"
+ "is_RRset A m ==> a \<in> A --> zgcd (a, m) = Numeral1"
apply (unfold is_RRset_def)
apply (rule RsetR.induct)
apply auto
@@ -181,7 +181,7 @@
lemma RsetR_zmult_mono:
"A \<in> RsetR m ==>
- #0 < m ==> zgcd (x, m) = #1 ==> (\<lambda>a. a * x) ` A \<in> RsetR m"
+ Numeral0 < m ==> zgcd (x, m) = Numeral1 ==> (\<lambda>a. a * x) ` A \<in> RsetR m"
apply (erule RsetR.induct)
apply simp_all
apply (rule RsetR.insert)
@@ -191,8 +191,8 @@
done
lemma card_nor_eq_noX:
- "#0 < m ==>
- zgcd (x, m) = #1 ==> card (noXRRset m x) = card (norRRset m)"
+ "Numeral0 < m ==>
+ zgcd (x, m) = Numeral1 ==> card (noXRRset m x) = card (norRRset m)"
apply (unfold norRRset_def noXRRset_def)
apply (rule card_image)
apply (auto simp add: inj_on_def Bnor_fin)
@@ -200,7 +200,7 @@
done
lemma noX_is_RRset:
- "#0 < m ==> zgcd (x, m) = #1 ==> is_RRset (noXRRset m x) m"
+ "Numeral0 < m ==> zgcd (x, m) = Numeral1 ==> is_RRset (noXRRset m x) m"
apply (unfold is_RRset_def phi_def)
apply (auto simp add: card_nor_eq_noX)
apply (unfold noXRRset_def norRRset_def)
@@ -210,7 +210,7 @@
done
lemma aux_some:
- "#1 < m ==> is_RRset A m ==> a \<in> A
+ "Numeral1 < m ==> is_RRset A m ==> a \<in> A
==> zcong a (SOME b. [a = b] (mod m) \<and> b \<in> norRRset m) m \<and>
(SOME b. [a = b] (mod m) \<and> b \<in> norRRset m) \<in> norRRset m"
apply (rule norR_mem_unique [THEN ex1_implies_ex, THEN someI_ex])
@@ -219,7 +219,7 @@
done
lemma RRset2norRR_correct:
- "#1 < m ==> is_RRset A m ==> a \<in> A ==>
+ "Numeral1 < m ==> is_RRset A m ==> a \<in> A ==>
[a = RRset2norRR A m a] (mod m) \<and> RRset2norRR A m a \<in> norRRset m"
apply (unfold RRset2norRR_def)
apply simp
@@ -238,7 +238,7 @@
done
lemma RRset_zcong_eq [rule_format]:
- "#1 < m ==>
+ "Numeral1 < m ==>
is_RRset A m ==> [a = b] (mod m) ==> a \<in> A --> b \<in> A --> a = b"
apply (unfold is_RRset_def)
apply (rule RsetR.induct)
@@ -252,7 +252,7 @@
done
lemma RRset2norRR_inj:
- "#1 < m ==> is_RRset A m ==> inj_on (RRset2norRR A m) A"
+ "Numeral1 < m ==> is_RRset A m ==> inj_on (RRset2norRR A m) A"
apply (unfold RRset2norRR_def inj_on_def)
apply auto
apply (subgoal_tac "\<exists>b. ([x = b] (mod m) \<and> b \<in> norRRset m) \<and>
@@ -267,7 +267,7 @@
done
lemma RRset2norRR_eq_norR:
- "#1 < m ==> is_RRset A m ==> RRset2norRR A m ` A = norRRset m"
+ "Numeral1 < m ==> is_RRset A m ==> RRset2norRR A m ` A = norRRset m"
apply (rule card_seteq)
prefer 3
apply (subst card_image)
@@ -286,7 +286,7 @@
done
lemma Bnor_prod_power [rule_format]:
- "x \<noteq> #0 ==> a < m --> setprod ((\<lambda>a. a * x) ` BnorRset (a, m)) =
+ "x \<noteq> Numeral0 ==> a < m --> setprod ((\<lambda>a. a * x) ` BnorRset (a, m)) =
setprod (BnorRset(a, m)) * x^card (BnorRset (a, m))"
apply (induct a m rule: BnorRset_induct)
prefer 2
@@ -313,7 +313,7 @@
done
lemma Bnor_prod_zgcd [rule_format]:
- "a < m --> zgcd (setprod (BnorRset (a, m)), m) = #1"
+ "a < m --> zgcd (setprod (BnorRset (a, m)), m) = Numeral1"
apply (induct a m rule: BnorRset_induct)
prefer 2
apply (subst BnorRset.simps)
@@ -324,12 +324,12 @@
done
theorem Euler_Fermat:
- "#0 < m ==> zgcd (x, m) = #1 ==> [x^(phi m) = #1] (mod m)"
+ "Numeral0 < m ==> zgcd (x, m) = Numeral1 ==> [x^(phi m) = Numeral1] (mod m)"
apply (unfold norRRset_def phi_def)
- apply (case_tac "x = #0")
- apply (case_tac [2] "m = #1")
+ apply (case_tac "x = Numeral0")
+ apply (case_tac [2] "m = Numeral1")
apply (rule_tac [3] iffD1)
- apply (rule_tac [3] k = "setprod (BnorRset (m - #1, m))"
+ apply (rule_tac [3] k = "setprod (BnorRset (m - Numeral1, m))"
in zcong_cancel2)
prefer 5
apply (subst Bnor_prod_power [symmetric])
@@ -352,7 +352,7 @@
lemma Bnor_prime [rule_format (no_asm)]:
"p \<in> zprime ==>
- a < p --> (\<forall>b. #0 < b \<and> b \<le> a --> zgcd (b, p) = #1)
+ a < p --> (\<forall>b. Numeral0 < b \<and> b \<le> a --> zgcd (b, p) = Numeral1)
--> card (BnorRset (a, p)) = nat a"
apply (unfold zprime_def)
apply (induct a p rule: BnorRset.induct)
@@ -361,7 +361,7 @@
apply auto
done
-lemma phi_prime: "p \<in> zprime ==> phi p = nat (p - #1)"
+lemma phi_prime: "p \<in> zprime ==> phi p = nat (p - Numeral1)"
apply (unfold phi_def norRRset_def)
apply (rule Bnor_prime)
apply auto
@@ -370,7 +370,7 @@
done
theorem Little_Fermat:
- "p \<in> zprime ==> \<not> p dvd x ==> [x^(nat (p - #1)) = #1] (mod p)"
+ "p \<in> zprime ==> \<not> p dvd x ==> [x^(nat (p - Numeral1)) = Numeral1] (mod p)"
apply (subst phi_prime [symmetric])
apply (rule_tac [2] Euler_Fermat)
apply (erule_tac [3] zprime_imp_zrelprime)
--- a/src/HOL/NumberTheory/Factorization.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/NumberTheory/Factorization.thy Fri Oct 05 21:52:39 2001 +0200
@@ -26,7 +26,7 @@
"nondec (x # xs) = (case xs of [] => True | y # ys => x \<le> y \<and> nondec xs)"
primrec
- "prod [] = 1'"
+ "prod [] = Suc 0"
"prod (x # xs) = x * prod xs"
primrec
@@ -40,12 +40,12 @@
subsection {* Arithmetic *}
-lemma one_less_m: "(m::nat) \<noteq> m * k ==> m \<noteq> 1' ==> 1' < m"
+lemma one_less_m: "(m::nat) \<noteq> m * k ==> m \<noteq> Suc 0 ==> Suc 0 < m"
apply (case_tac m)
apply auto
done
-lemma one_less_k: "(m::nat) \<noteq> m * k ==> 1' < m * k ==> 1' < k"
+lemma one_less_k: "(m::nat) \<noteq> m * k ==> Suc 0 < m * k ==> Suc 0 < k"
apply (case_tac k)
apply auto
done
@@ -54,13 +54,13 @@
apply auto
done
-lemma mn_eq_m_one: "(0::nat) < m ==> m * n = m ==> n = 1'"
+lemma mn_eq_m_one: "(0::nat) < m ==> m * n = m ==> n = Suc 0"
apply (case_tac n)
apply auto
done
lemma prod_mn_less_k:
- "(0::nat) < n ==> 0 < k ==> 1' < m ==> m * n = k ==> n < k"
+ "(0::nat) < n ==> 0 < k ==> Suc 0 < m ==> m * n = k ==> n < k"
apply (induct m)
apply auto
done
@@ -88,7 +88,7 @@
apply auto
done
-lemma prime_nd_one: "p \<in> prime ==> \<not> p dvd 1'"
+lemma prime_nd_one: "p \<in> prime ==> \<not> p dvd Suc 0"
apply (unfold prime_def dvd_def)
apply auto
done
@@ -115,13 +115,13 @@
apply auto
done
-lemma primel_one_empty: "primel xs ==> prod xs = 1' ==> xs = []"
+lemma primel_one_empty: "primel xs ==> prod xs = Suc 0 ==> xs = []"
apply (unfold primel_def prime_def)
apply (case_tac xs)
apply simp_all
done
-lemma prime_g_one: "p \<in> prime ==> 1' < p"
+lemma prime_g_one: "p \<in> prime ==> Suc 0 < p"
apply (unfold prime_def)
apply auto
done
@@ -132,7 +132,7 @@
done
lemma primel_nempty_g_one [rule_format]:
- "primel xs --> xs \<noteq> [] --> 1' < prod xs"
+ "primel xs --> xs \<noteq> [] --> Suc 0 < prod xs"
apply (unfold primel_def prime_def)
apply (induct xs)
apply (auto elim: one_less_mult)
@@ -223,8 +223,8 @@
done
lemma not_prime_ex_mk:
- "1' < n \<and> n \<notin> prime ==>
- \<exists>m k. 1' < m \<and> 1' < k \<and> m < n \<and> k < n \<and> n = m * k"
+ "Suc 0 < n \<and> n \<notin> prime ==>
+ \<exists>m k. Suc 0 < m \<and> Suc 0 < k \<and> m < n \<and> k < n \<and> n = m * k"
apply (unfold prime_def dvd_def)
apply (auto intro: n_less_m_mult_n n_less_n_mult_m one_less_m one_less_k)
done
@@ -237,7 +237,7 @@
apply (simp add: primel_append)
done
-lemma factor_exists [rule_format]: "1' < n --> (\<exists>l. primel l \<and> prod l = n)"
+lemma factor_exists [rule_format]: "Suc 0 < n --> (\<exists>l. primel l \<and> prod l = n)"
apply (induct n rule: nat_less_induct)
apply (rule impI)
apply (case_tac "n \<in> prime")
@@ -247,7 +247,7 @@
apply (auto intro!: split_primel)
done
-lemma nondec_factor_exists: "1' < n ==> \<exists>l. primel l \<and> nondec l \<and> prod l = n"
+lemma nondec_factor_exists: "Suc 0 < n ==> \<exists>l. primel l \<and> nondec l \<and> prod l = n"
apply (erule factor_exists [THEN exE])
apply (blast intro!: ex_nondec_lemma)
done
@@ -349,7 +349,7 @@
done
lemma unique_prime_factorization [rule_format]:
- "\<forall>n. 1' < n --> (\<exists>!l. primel l \<and> nondec l \<and> prod l = n)"
+ "\<forall>n. Suc 0 < n --> (\<exists>!l. primel l \<and> nondec l \<and> prod l = n)"
apply safe
apply (erule nondec_factor_exists)
apply (rule perm_nondec_unique)
--- a/src/HOL/NumberTheory/Fib.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/NumberTheory/Fib.thy Fri Oct 05 21:52:39 2001 +0200
@@ -19,7 +19,7 @@
consts fib :: "nat => nat"
recdef fib less_than
zero: "fib 0 = 0"
- one: "fib 1' = 1'"
+ one: "fib (Suc 0) = Suc 0"
Suc_Suc: "fib (Suc (Suc x)) = fib x + fib (Suc x)"
text {*
@@ -67,21 +67,21 @@
*}
lemma fib_Cassini: "int (fib (Suc (Suc n)) * fib n) =
- (if n mod #2 = 0 then int (fib (Suc n) * fib (Suc n)) - #1
- else int (fib (Suc n) * fib (Suc n)) + #1)"
+ (if n mod # 2 = 0 then int (fib (Suc n) * fib (Suc n)) - Numeral1
+ else int (fib (Suc n) * fib (Suc n)) + Numeral1)"
apply (induct n rule: fib.induct)
apply (simp add: fib.Suc_Suc)
apply (simp add: fib.Suc_Suc mod_Suc)
apply (simp add: fib.Suc_Suc
add_mult_distrib add_mult_distrib2 mod_Suc zmult_int [symmetric] zmult_ac)
- apply (subgoal_tac "x mod #2 < #2", arith)
+ apply (subgoal_tac "x mod # 2 < # 2", arith)
apply simp
done
text {* \medskip Towards Law 6.111 of Concrete Mathematics *}
-lemma gcd_fib_Suc_eq_1: "gcd (fib n, fib (Suc n)) = 1'"
+lemma gcd_fib_Suc_eq_1: "gcd (fib n, fib (Suc n)) = Suc 0"
apply (induct n rule: fib.induct)
prefer 3
apply (simp add: gcd_commute fib_Suc3)
--- a/src/HOL/NumberTheory/IntFact.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/NumberTheory/IntFact.thy Fri Oct 05 21:52:39 2001 +0200
@@ -10,7 +10,7 @@
text {*
Factorial on integers and recursively defined set including all
- Integers from @{term 2} up to @{term a}. Plus definition of product
+ Integers from @{text 2} up to @{text a}. Plus definition of product
of finite set.
\bigskip
@@ -22,18 +22,18 @@
d22set :: "int => int set"
recdef zfact "measure ((\<lambda>n. nat n) :: int => nat)"
- "zfact n = (if n \<le> #0 then #1 else n * zfact (n - #1))"
+ "zfact n = (if n \<le> Numeral0 then Numeral1 else n * zfact (n - Numeral1))"
defs
- setprod_def: "setprod A == (if finite A then fold (op *) #1 A else #1)"
+ setprod_def: "setprod A == (if finite A then fold (op *) Numeral1 A else Numeral1)"
recdef d22set "measure ((\<lambda>a. nat a) :: int => nat)"
- "d22set a = (if #1 < a then insert a (d22set (a - #1)) else {})"
+ "d22set a = (if Numeral1 < a then insert a (d22set (a - Numeral1)) else {})"
text {* \medskip @{term setprod} --- product of finite set *}
-lemma setprod_empty [simp]: "setprod {} = #1"
+lemma setprod_empty [simp]: "setprod {} = Numeral1"
apply (simp add: setprod_def)
done
@@ -46,7 +46,7 @@
text {*
\medskip @{term d22set} --- recursively defined set including all
- integers from @{term 2} up to @{term a}
+ integers from @{text 2} up to @{text a}
*}
declare d22set.simps [simp del]
@@ -54,7 +54,7 @@
lemma d22set_induct:
"(!!a. P {} a) ==>
- (!!a. #1 < (a::int) ==> P (d22set (a - #1)) (a - #1)
+ (!!a. Numeral1 < (a::int) ==> P (d22set (a - Numeral1)) (a - Numeral1)
==> P (d22set a) a)
==> P (d22set u) u"
proof -
@@ -62,14 +62,14 @@
show ?thesis
apply (rule d22set.induct)
apply safe
- apply (case_tac [2] "#1 < a")
+ apply (case_tac [2] "Numeral1 < a")
apply (rule_tac [2] rule_context)
apply (simp_all (no_asm_simp))
apply (simp_all (no_asm_simp) add: d22set.simps rule_context)
done
qed
-lemma d22set_g_1 [rule_format]: "b \<in> d22set a --> #1 < b"
+lemma d22set_g_1 [rule_format]: "b \<in> d22set a --> Numeral1 < b"
apply (induct a rule: d22set_induct)
prefer 2
apply (subst d22set.simps)
@@ -87,7 +87,7 @@
apply (auto dest: d22set_le)
done
-lemma d22set_mem [rule_format]: "#1 < b --> b \<le> a --> b \<in> d22set a"
+lemma d22set_mem [rule_format]: "Numeral1 < b --> b \<le> a --> b \<in> d22set a"
apply (induct a rule: d22set.induct)
apply auto
apply (simp_all add: d22set.simps)
@@ -109,7 +109,7 @@
apply (simp add: d22set.simps zfact.simps)
apply (subst d22set.simps)
apply (subst zfact.simps)
- apply (case_tac "#1 < a")
+ apply (case_tac "Numeral1 < a")
prefer 2
apply (simp add: d22set.simps zfact.simps)
apply (simp add: d22set_fin d22set_le_swap)
--- a/src/HOL/NumberTheory/IntPrimes.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/NumberTheory/IntPrimes.thy Fri Oct 05 21:52:39 2001 +0200
@@ -29,7 +29,7 @@
"measure ((\<lambda>(m, n, r', r, s', s, t', t). nat r)
:: int * int * int * int *int * int * int * int => nat)"
"xzgcda (m, n, r', r, s', s, t', t) =
- (if r \<le> #0 then (r', s', t')
+ (if r \<le> Numeral0 then (r', s', t')
else xzgcda (m, n, r, r' mod r, s, s' - (r' div r) * s, t, t' - (r' div r) * t))"
(hints simp: pos_mod_bound)
@@ -38,13 +38,13 @@
"zgcd == \<lambda>(x,y). int (gcd (nat (abs x), nat (abs y)))"
defs
- xzgcd_def: "xzgcd m n == xzgcda (m, n, m, n, #1, #0, #0, #1)"
- zprime_def: "zprime == {p. #1 < p \<and> (\<forall>m. m dvd p --> m = #1 \<or> m = p)}"
+ xzgcd_def: "xzgcd m n == xzgcda (m, n, m, n, Numeral1, Numeral0, Numeral0, Numeral1)"
+ zprime_def: "zprime == {p. Numeral1 < p \<and> (\<forall>m. m dvd p --> m = Numeral1 \<or> m = p)}"
zcong_def: "[a = b] (mod m) == m dvd (a - b)"
lemma zabs_eq_iff:
- "(abs (z::int) = w) = (z = w \<and> #0 <= z \<or> z = -w \<and> z < #0)"
+ "(abs (z::int) = w) = (z = w \<and> Numeral0 <= z \<or> z = -w \<and> z < Numeral0)"
apply (auto simp add: zabs_def)
done
@@ -64,17 +64,17 @@
subsection {* Divides relation *}
-lemma zdvd_0_right [iff]: "(m::int) dvd #0"
+lemma zdvd_0_right [iff]: "(m::int) dvd Numeral0"
apply (unfold dvd_def)
apply (blast intro: zmult_0_right [symmetric])
done
-lemma zdvd_0_left [iff]: "(#0 dvd (m::int)) = (m = #0)"
+lemma zdvd_0_left [iff]: "(Numeral0 dvd (m::int)) = (m = Numeral0)"
apply (unfold dvd_def)
apply auto
done
-lemma zdvd_1_left [iff]: "#1 dvd (m::int)"
+lemma zdvd_1_left [iff]: "Numeral1 dvd (m::int)"
apply (unfold dvd_def)
apply simp
done
@@ -104,7 +104,7 @@
done
lemma zdvd_anti_sym:
- "#0 < m ==> #0 < n ==> m dvd n ==> n dvd m ==> m = (n::int)"
+ "Numeral0 < m ==> Numeral0 < n ==> m dvd n ==> n dvd m ==> m = (n::int)"
apply (unfold dvd_def)
apply auto
apply (simp add: zmult_assoc zmult_eq_self_iff int_0_less_mult_iff zmult_eq_1_iff)
@@ -186,19 +186,19 @@
apply (simp add: zdvd_zadd zdvd_zmult2)
done
-lemma zdvd_iff_zmod_eq_0: "(k dvd n) = (n mod (k::int) = #0)"
+lemma zdvd_iff_zmod_eq_0: "(k dvd n) = (n mod (k::int) = Numeral0)"
apply (unfold dvd_def)
apply auto
done
-lemma zdvd_not_zless: "#0 < m ==> m < n ==> \<not> n dvd (m::int)"
+lemma zdvd_not_zless: "Numeral0 < m ==> m < n ==> \<not> n dvd (m::int)"
apply (unfold dvd_def)
apply auto
- apply (subgoal_tac "#0 < n")
+ apply (subgoal_tac "Numeral0 < n")
prefer 2
apply (blast intro: zless_trans)
apply (simp add: int_0_less_mult_iff)
- apply (subgoal_tac "n * k < n * #1")
+ apply (subgoal_tac "n * k < n * Numeral1")
apply (drule zmult_zless_cancel1 [THEN iffD1])
apply auto
done
@@ -221,7 +221,7 @@
nat_mult_distrib [symmetric] nat_eq_iff2)
done
-lemma nat_dvd_iff: "(nat z dvd m) = (if #0 \<le> z then (z dvd int m) else m = 0)"
+lemma nat_dvd_iff: "(nat z dvd m) = (if Numeral0 \<le> z then (z dvd int m) else m = 0)"
apply (auto simp add: dvd_def zmult_int [symmetric])
apply (rule_tac x = "nat k" in exI)
apply (cut_tac k = m in int_less_0_conv)
@@ -245,11 +245,11 @@
subsection {* Euclid's Algorithm and GCD *}
-lemma zgcd_0 [simp]: "zgcd (m, #0) = abs m"
+lemma zgcd_0 [simp]: "zgcd (m, Numeral0) = abs m"
apply (simp add: zgcd_def zabs_def)
done
-lemma zgcd_0_left [simp]: "zgcd (#0, m) = abs m"
+lemma zgcd_0_left [simp]: "zgcd (Numeral0, m) = abs m"
apply (simp add: zgcd_def zabs_def)
done
@@ -261,7 +261,7 @@
apply (simp add: zgcd_def)
done
-lemma zgcd_non_0: "#0 < n ==> zgcd (m, n) = zgcd (n, m mod n)"
+lemma zgcd_non_0: "Numeral0 < n ==> zgcd (m, n) = zgcd (n, m mod n)"
apply (frule_tac b = n and a = m in pos_mod_sign)
apply (simp add: zgcd_def zabs_def nat_mod_distrib)
apply (cut_tac a = "-m" and b = n in zmod_zminus1_eq_if)
@@ -273,17 +273,17 @@
done
lemma zgcd_eq: "zgcd (m, n) = zgcd (n, m mod n)"
- apply (tactic {* zdiv_undefined_case_tac "n = #0" 1 *})
+ apply (tactic {* zdiv_undefined_case_tac "n = Numeral0" 1 *})
apply (auto simp add: linorder_neq_iff zgcd_non_0)
apply (cut_tac m = "-m" and n = "-n" in zgcd_non_0)
apply auto
done
-lemma zgcd_1 [simp]: "zgcd (m, #1) = #1"
+lemma zgcd_1 [simp]: "zgcd (m, Numeral1) = Numeral1"
apply (simp add: zgcd_def zabs_def)
done
-lemma zgcd_0_1_iff [simp]: "(zgcd (#0, m) = #1) = (abs m = #1)"
+lemma zgcd_0_1_iff [simp]: "(zgcd (Numeral0, m) = Numeral1) = (abs m = Numeral1)"
apply (simp add: zgcd_def zabs_def)
done
@@ -303,7 +303,7 @@
apply (simp add: zgcd_def gcd_commute)
done
-lemma zgcd_1_left [simp]: "zgcd (#1, m) = #1"
+lemma zgcd_1_left [simp]: "zgcd (Numeral1, m) = Numeral1"
apply (simp add: zgcd_def gcd_1_left)
done
@@ -320,7 +320,7 @@
lemmas zgcd_ac = zgcd_assoc zgcd_commute zgcd_left_commute
-- {* addition is an AC-operator *}
-lemma zgcd_zmult_distrib2: "#0 \<le> k ==> k * zgcd (m, n) = zgcd (k * m, k * n)"
+lemma zgcd_zmult_distrib2: "Numeral0 \<le> k ==> k * zgcd (m, n) = zgcd (k * m, k * n)"
apply (simp del: zmult_zminus_right
add: zmult_zminus_right [symmetric] nat_mult_distrib zgcd_def zabs_def
zmult_less_0_iff gcd_mult_distrib2 [symmetric] zmult_int [symmetric])
@@ -330,29 +330,29 @@
apply (simp add: zabs_def zgcd_zmult_distrib2)
done
-lemma zgcd_self [simp]: "#0 \<le> m ==> zgcd (m, m) = m"
- apply (cut_tac k = m and m = "#1" and n = "#1" in zgcd_zmult_distrib2)
+lemma zgcd_self [simp]: "Numeral0 \<le> m ==> zgcd (m, m) = m"
+ apply (cut_tac k = m and m = "Numeral1" and n = "Numeral1" in zgcd_zmult_distrib2)
apply simp_all
done
-lemma zgcd_zmult_eq_self [simp]: "#0 \<le> k ==> zgcd (k, k * n) = k"
- apply (cut_tac k = k and m = "#1" and n = n in zgcd_zmult_distrib2)
+lemma zgcd_zmult_eq_self [simp]: "Numeral0 \<le> k ==> zgcd (k, k * n) = k"
+ apply (cut_tac k = k and m = "Numeral1" and n = n in zgcd_zmult_distrib2)
apply simp_all
done
-lemma zgcd_zmult_eq_self2 [simp]: "#0 \<le> k ==> zgcd (k * n, k) = k"
- apply (cut_tac k = k and m = n and n = "#1" in zgcd_zmult_distrib2)
+lemma zgcd_zmult_eq_self2 [simp]: "Numeral0 \<le> k ==> zgcd (k * n, k) = k"
+ apply (cut_tac k = k and m = n and n = "Numeral1" in zgcd_zmult_distrib2)
apply simp_all
done
-lemma aux: "zgcd (n, k) = #1 ==> k dvd m * n ==> #0 \<le> m ==> k dvd m"
+lemma aux: "zgcd (n, k) = Numeral1 ==> k dvd m * n ==> Numeral0 \<le> m ==> k dvd m"
apply (subgoal_tac "m = zgcd (m * n, m * k)")
apply (erule ssubst, rule zgcd_greatest_iff [THEN iffD2])
apply (simp_all add: zgcd_zmult_distrib2 [symmetric] int_0_le_mult_iff)
done
-lemma zrelprime_zdvd_zmult: "zgcd (n, k) = #1 ==> k dvd m * n ==> k dvd m"
- apply (case_tac "#0 \<le> m")
+lemma zrelprime_zdvd_zmult: "zgcd (n, k) = Numeral1 ==> k dvd m * n ==> k dvd m"
+ apply (case_tac "Numeral0 \<le> m")
apply (blast intro: aux)
apply (subgoal_tac "k dvd -m")
apply (rule_tac [2] aux)
@@ -360,20 +360,20 @@
done
lemma zprime_imp_zrelprime:
- "p \<in> zprime ==> \<not> p dvd n ==> zgcd (n, p) = #1"
+ "p \<in> zprime ==> \<not> p dvd n ==> zgcd (n, p) = Numeral1"
apply (unfold zprime_def)
apply auto
done
lemma zless_zprime_imp_zrelprime:
- "p \<in> zprime ==> #0 < n ==> n < p ==> zgcd (n, p) = #1"
+ "p \<in> zprime ==> Numeral0 < n ==> n < p ==> zgcd (n, p) = Numeral1"
apply (erule zprime_imp_zrelprime)
apply (erule zdvd_not_zless)
apply assumption
done
lemma zprime_zdvd_zmult:
- "#0 \<le> (m::int) ==> p \<in> zprime ==> p dvd m * n ==> p dvd m \<or> p dvd n"
+ "Numeral0 \<le> (m::int) ==> p \<in> zprime ==> p dvd m * n ==> p dvd m \<or> p dvd n"
apply safe
apply (rule zrelprime_zdvd_zmult)
apply (rule zprime_imp_zrelprime)
@@ -392,7 +392,7 @@
done
lemma zgcd_zmult_zdvd_zgcd:
- "zgcd (k, n) = #1 ==> zgcd (k * m, n) dvd zgcd (m, n)"
+ "zgcd (k, n) = Numeral1 ==> zgcd (k * m, n) dvd zgcd (m, n)"
apply (simp add: zgcd_greatest_iff)
apply (rule_tac n = k in zrelprime_zdvd_zmult)
prefer 2
@@ -402,16 +402,16 @@
apply (simp (no_asm) add: zgcd_ac)
done
-lemma zgcd_zmult_cancel: "zgcd (k, n) = #1 ==> zgcd (k * m, n) = zgcd (m, n)"
+lemma zgcd_zmult_cancel: "zgcd (k, n) = Numeral1 ==> zgcd (k * m, n) = zgcd (m, n)"
apply (simp add: zgcd_def nat_abs_mult_distrib gcd_mult_cancel)
done
lemma zgcd_zgcd_zmult:
- "zgcd (k, m) = #1 ==> zgcd (n, m) = #1 ==> zgcd (k * n, m) = #1"
+ "zgcd (k, m) = Numeral1 ==> zgcd (n, m) = Numeral1 ==> zgcd (k * n, m) = Numeral1"
apply (simp (no_asm_simp) add: zgcd_zmult_cancel)
done
-lemma zdvd_iff_zgcd: "#0 < m ==> (m dvd n) = (zgcd (n, m) = m)"
+lemma zdvd_iff_zgcd: "Numeral0 < m ==> (m dvd n) = (zgcd (n, m) = m)"
apply safe
apply (rule_tac [2] n = "zgcd (n, m)" in zdvd_trans)
apply (rule_tac [3] zgcd_zdvd1)
@@ -423,7 +423,7 @@
subsection {* Congruences *}
-lemma zcong_1 [simp]: "[a = b] (mod #1)"
+lemma zcong_1 [simp]: "[a = b] (mod Numeral1)"
apply (unfold zcong_def)
apply auto
done
@@ -494,19 +494,19 @@
done
lemma zcong_square:
- "p \<in> zprime ==> #0 < a ==> [a * a = #1] (mod p)
- ==> [a = #1] (mod p) \<or> [a = p - #1] (mod p)"
+ "p \<in> zprime ==> Numeral0 < a ==> [a * a = Numeral1] (mod p)
+ ==> [a = Numeral1] (mod p) \<or> [a = p - Numeral1] (mod p)"
apply (unfold zcong_def)
apply (rule zprime_zdvd_zmult)
- apply (rule_tac [3] s = "a * a - #1 + p * (#1 - a)" in subst)
+ apply (rule_tac [3] s = "a * a - Numeral1 + p * (Numeral1 - a)" in subst)
prefer 4
apply (simp add: zdvd_reduce)
apply (simp_all add: zdiff_zmult_distrib zmult_commute zdiff_zmult_distrib2)
done
lemma zcong_cancel:
- "#0 \<le> m ==>
- zgcd (k, m) = #1 ==> [a * k = b * k] (mod m) = [a = b] (mod m)"
+ "Numeral0 \<le> m ==>
+ zgcd (k, m) = Numeral1 ==> [a * k = b * k] (mod m) = [a = b] (mod m)"
apply safe
prefer 2
apply (blast intro: zcong_scalar)
@@ -523,19 +523,19 @@
done
lemma zcong_cancel2:
- "#0 \<le> m ==>
- zgcd (k, m) = #1 ==> [k * a = k * b] (mod m) = [a = b] (mod m)"
+ "Numeral0 \<le> m ==>
+ zgcd (k, m) = Numeral1 ==> [k * a = k * b] (mod m) = [a = b] (mod m)"
apply (simp add: zmult_commute zcong_cancel)
done
lemma zcong_zgcd_zmult_zmod:
- "[a = b] (mod m) ==> [a = b] (mod n) ==> zgcd (m, n) = #1
+ "[a = b] (mod m) ==> [a = b] (mod n) ==> zgcd (m, n) = Numeral1
==> [a = b] (mod m * n)"
apply (unfold zcong_def dvd_def)
apply auto
apply (subgoal_tac "m dvd n * ka")
apply (subgoal_tac "m dvd ka")
- apply (case_tac [2] "#0 \<le> ka")
+ apply (case_tac [2] "Numeral0 \<le> ka")
prefer 3
apply (subst zdvd_zminus_iff [symmetric])
apply (rule_tac n = n in zrelprime_zdvd_zmult)
@@ -550,8 +550,8 @@
done
lemma zcong_zless_imp_eq:
- "#0 \<le> a ==>
- a < m ==> #0 \<le> b ==> b < m ==> [a = b] (mod m) ==> a = b"
+ "Numeral0 \<le> a ==>
+ a < m ==> Numeral0 \<le> b ==> b < m ==> [a = b] (mod m) ==> a = b"
apply (unfold zcong_def dvd_def)
apply auto
apply (drule_tac f = "\<lambda>z. z mod m" in arg_cong)
@@ -566,38 +566,38 @@
done
lemma zcong_square_zless:
- "p \<in> zprime ==> #0 < a ==> a < p ==>
- [a * a = #1] (mod p) ==> a = #1 \<or> a = p - #1"
+ "p \<in> zprime ==> Numeral0 < a ==> a < p ==>
+ [a * a = Numeral1] (mod p) ==> a = Numeral1 \<or> a = p - Numeral1"
apply (cut_tac p = p and a = a in zcong_square)
apply (simp add: zprime_def)
apply (auto intro: zcong_zless_imp_eq)
done
lemma zcong_not:
- "#0 < a ==> a < m ==> #0 < b ==> b < a ==> \<not> [a = b] (mod m)"
+ "Numeral0 < a ==> a < m ==> Numeral0 < b ==> b < a ==> \<not> [a = b] (mod m)"
apply (unfold zcong_def)
apply (rule zdvd_not_zless)
apply auto
done
lemma zcong_zless_0:
- "#0 \<le> a ==> a < m ==> [a = #0] (mod m) ==> a = #0"
+ "Numeral0 \<le> a ==> a < m ==> [a = Numeral0] (mod m) ==> a = Numeral0"
apply (unfold zcong_def dvd_def)
apply auto
- apply (subgoal_tac "#0 < m")
+ apply (subgoal_tac "Numeral0 < m")
apply (rotate_tac -1)
apply (simp add: int_0_le_mult_iff)
- apply (subgoal_tac "m * k < m * #1")
+ apply (subgoal_tac "m * k < m * Numeral1")
apply (drule zmult_zless_cancel1 [THEN iffD1])
apply (auto simp add: linorder_neq_iff)
done
lemma zcong_zless_unique:
- "#0 < m ==> (\<exists>!b. #0 \<le> b \<and> b < m \<and> [a = b] (mod m))"
+ "Numeral0 < m ==> (\<exists>!b. Numeral0 \<le> b \<and> b < m \<and> [a = b] (mod m))"
apply auto
apply (subgoal_tac [2] "[b = y] (mod m)")
- apply (case_tac [2] "b = #0")
- apply (case_tac [3] "y = #0")
+ apply (case_tac [2] "b = Numeral0")
+ apply (case_tac [3] "y = Numeral0")
apply (auto intro: zcong_trans zcong_zless_0 zcong_zless_imp_eq order_less_le
simp add: zcong_sym)
apply (unfold zcong_def dvd_def)
@@ -616,8 +616,8 @@
done
lemma zgcd_zcong_zgcd:
- "#0 < m ==>
- zgcd (a, m) = #1 ==> [a = b] (mod m) ==> zgcd (b, m) = #1"
+ "Numeral0 < m ==>
+ zgcd (a, m) = Numeral1 ==> [a = b] (mod m) ==> zgcd (b, m) = Numeral1"
apply (auto simp add: zcong_iff_lin)
done
@@ -643,7 +643,7 @@
apply (simp add: zadd_commute)
done
-lemma zcong_zmod_eq: "#0 < m ==> [a = b] (mod m) = (a mod m = b mod m)"
+lemma zcong_zmod_eq: "Numeral0 < m ==> [a = b] (mod m) = (a mod m = b mod m)"
apply auto
apply (rule_tac m = m in zcong_zless_imp_eq)
prefer 5
@@ -659,13 +659,13 @@
apply (auto simp add: zcong_def)
done
-lemma zcong_zero [iff]: "[a = b] (mod #0) = (a = b)"
+lemma zcong_zero [iff]: "[a = b] (mod Numeral0) = (a = b)"
apply (auto simp add: zcong_def)
done
lemma "[a = b] (mod m) = (a mod m = b mod m)"
- apply (tactic {* zdiv_undefined_case_tac "m = #0" 1 *})
- apply (case_tac "#0 < m")
+ apply (tactic {* zdiv_undefined_case_tac "m = Numeral0" 1 *})
+ apply (case_tac "Numeral0 < m")
apply (simp add: zcong_zmod_eq)
apply (rule_tac t = m in zminus_zminus [THEN subst])
apply (subst zcong_zminus)
@@ -677,7 +677,7 @@
subsection {* Modulo *}
lemma zmod_zdvd_zmod:
- "#0 < (m::int) ==> m dvd b ==> (a mod b mod m) = (a mod m)"
+ "Numeral0 < (m::int) ==> m dvd b ==> (a mod b mod m) = (a mod m)"
apply (unfold dvd_def)
apply auto
apply (subst zcong_zmod_eq [symmetric])
@@ -696,14 +696,14 @@
declare xzgcda.simps [simp del]
lemma aux1:
- "zgcd (r', r) = k --> #0 < r -->
+ "zgcd (r', r) = k --> Numeral0 < r -->
(\<exists>sn tn. xzgcda (m, n, r', r, s', s, t', t) = (k, sn, tn))"
apply (rule_tac u = m and v = n and w = r' and x = r and y = s' and
z = s and aa = t' and ab = t in xzgcda.induct)
apply (subst zgcd_eq)
apply (subst xzgcda.simps)
apply auto
- apply (case_tac "r' mod r = #0")
+ apply (case_tac "r' mod r = Numeral0")
prefer 2
apply (frule_tac a = "r'" in pos_mod_sign)
apply auto
@@ -716,14 +716,14 @@
done
lemma aux2:
- "(\<exists>sn tn. xzgcda (m, n, r', r, s', s, t', t) = (k, sn, tn)) --> #0 < r -->
+ "(\<exists>sn tn. xzgcda (m, n, r', r, s', s, t', t) = (k, sn, tn)) --> Numeral0 < r -->
zgcd (r', r) = k"
apply (rule_tac u = m and v = n and w = r' and x = r and y = s' and
z = s and aa = t' and ab = t in xzgcda.induct)
apply (subst zgcd_eq)
apply (subst xzgcda.simps)
apply (auto simp add: linorder_not_le)
- apply (case_tac "r' mod r = #0")
+ apply (case_tac "r' mod r = Numeral0")
prefer 2
apply (frule_tac a = "r'" in pos_mod_sign)
apply auto
@@ -735,7 +735,7 @@
done
lemma xzgcd_correct:
- "#0 < n ==> (zgcd (m, n) = k) = (\<exists>s t. xzgcd m n = (k, s, t))"
+ "Numeral0 < n ==> (zgcd (m, n) = k) = (\<exists>s t. xzgcd m n = (k, s, t))"
apply (unfold xzgcd_def)
apply (rule iffI)
apply (rule_tac [2] aux2 [THEN mp, THEN mp])
@@ -768,17 +768,17 @@
by (rule iffD2 [OF order_less_le conjI])
lemma xzgcda_linear [rule_format]:
- "#0 < r --> xzgcda (m, n, r', r, s', s, t', t) = (rn, sn, tn) -->
+ "Numeral0 < r --> xzgcda (m, n, r', r, s', s, t', t) = (rn, sn, tn) -->
r' = s' * m + t' * n --> r = s * m + t * n --> rn = sn * m + tn * n"
apply (rule_tac u = m and v = n and w = r' and x = r and y = s' and
z = s and aa = t' and ab = t in xzgcda.induct)
apply (subst xzgcda.simps)
apply (simp (no_asm))
apply (rule impI)+
- apply (case_tac "r' mod r = #0")
+ apply (case_tac "r' mod r = Numeral0")
apply (simp add: xzgcda.simps)
apply clarify
- apply (subgoal_tac "#0 < r' mod r")
+ apply (subgoal_tac "Numeral0 < r' mod r")
apply (rule_tac [2] order_le_neq_implies_less)
apply (rule_tac [2] pos_mod_sign)
apply (cut_tac m = m and n = n and r' = r' and r = r and s' = s' and
@@ -787,7 +787,7 @@
done
lemma xzgcd_linear:
- "#0 < n ==> xzgcd m n = (r, s, t) ==> r = s * m + t * n"
+ "Numeral0 < n ==> xzgcd m n = (r, s, t) ==> r = s * m + t * n"
apply (unfold xzgcd_def)
apply (erule xzgcda_linear)
apply assumption
@@ -795,7 +795,7 @@
done
lemma zgcd_ex_linear:
- "#0 < n ==> zgcd (m, n) = k ==> (\<exists>s t. k = s * m + t * n)"
+ "Numeral0 < n ==> zgcd (m, n) = k ==> (\<exists>s t. k = s * m + t * n)"
apply (simp add: xzgcd_correct)
apply safe
apply (rule exI)+
@@ -804,8 +804,8 @@
done
lemma zcong_lineq_ex:
- "#0 < n ==> zgcd (a, n) = #1 ==> \<exists>x. [a * x = #1] (mod n)"
- apply (cut_tac m = a and n = n and k = "#1" in zgcd_ex_linear)
+ "Numeral0 < n ==> zgcd (a, n) = Numeral1 ==> \<exists>x. [a * x = Numeral1] (mod n)"
+ apply (cut_tac m = a and n = n and k = "Numeral1" in zgcd_ex_linear)
apply safe
apply (rule_tac x = s in exI)
apply (rule_tac b = "s * a + t * n" in zcong_trans)
@@ -816,8 +816,8 @@
done
lemma zcong_lineq_unique:
- "#0 < n ==>
- zgcd (a, n) = #1 ==> \<exists>!x. #0 \<le> x \<and> x < n \<and> [a * x = b] (mod n)"
+ "Numeral0 < n ==>
+ zgcd (a, n) = Numeral1 ==> \<exists>!x. Numeral0 \<le> x \<and> x < n \<and> [a * x = b] (mod n)"
apply auto
apply (rule_tac [2] zcong_zless_imp_eq)
apply (tactic {* stac (thm "zcong_cancel2" RS sym) 6 *})
@@ -833,7 +833,7 @@
apply (subst zcong_zmod)
apply (subst zmod_zmult1_eq [symmetric])
apply (subst zcong_zmod [symmetric])
- apply (subgoal_tac "[a * x * b = #1 * b] (mod n)")
+ apply (subgoal_tac "[a * x * b = Numeral1 * b] (mod n)")
apply (rule_tac [2] zcong_zmult)
apply (simp_all add: zmult_assoc)
done
--- a/src/HOL/NumberTheory/WilsonBij.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/NumberTheory/WilsonBij.thy Fri Oct 05 21:52:39 2001 +0200
@@ -20,19 +20,19 @@
constdefs
reciR :: "int => int => int => bool"
"reciR p ==
- \<lambda>a b. zcong (a * b) #1 p \<and> #1 < a \<and> a < p - #1 \<and> #1 < b \<and> b < p - #1"
+ \<lambda>a b. zcong (a * b) Numeral1 p \<and> Numeral1 < a \<and> a < p - Numeral1 \<and> Numeral1 < b \<and> b < p - Numeral1"
inv :: "int => int => int"
"inv p a ==
- if p \<in> zprime \<and> #0 < a \<and> a < p then
- (SOME x. #0 \<le> x \<and> x < p \<and> zcong (a * x) #1 p)
- else #0"
+ if p \<in> zprime \<and> Numeral0 < a \<and> a < p then
+ (SOME x. Numeral0 \<le> x \<and> x < p \<and> zcong (a * x) Numeral1 p)
+ else Numeral0"
text {* \medskip Inverse *}
lemma inv_correct:
- "p \<in> zprime ==> #0 < a ==> a < p
- ==> #0 \<le> inv p a \<and> inv p a < p \<and> [a * inv p a = #1] (mod p)"
+ "p \<in> zprime ==> Numeral0 < a ==> a < p
+ ==> Numeral0 \<le> inv p a \<and> inv p a < p \<and> [a * inv p a = Numeral1] (mod p)"
apply (unfold inv_def)
apply (simp (no_asm_simp))
apply (rule zcong_lineq_unique [THEN ex1_implies_ex, THEN someI_ex])
@@ -46,53 +46,53 @@
lemmas inv_is_inv = inv_correct [THEN conjunct2, THEN conjunct2, standard]
lemma inv_not_0:
- "p \<in> zprime ==> #1 < a ==> a < p - #1 ==> inv p a \<noteq> #0"
+ "p \<in> zprime ==> Numeral1 < a ==> a < p - Numeral1 ==> inv p a \<noteq> Numeral0"
-- {* same as @{text WilsonRuss} *}
apply safe
apply (cut_tac a = a and p = p in inv_is_inv)
apply (unfold zcong_def)
apply auto
- apply (subgoal_tac "\<not> p dvd #1")
+ apply (subgoal_tac "\<not> p dvd Numeral1")
apply (rule_tac [2] zdvd_not_zless)
- apply (subgoal_tac "p dvd #1")
+ apply (subgoal_tac "p dvd Numeral1")
prefer 2
apply (subst zdvd_zminus_iff [symmetric])
apply auto
done
lemma inv_not_1:
- "p \<in> zprime ==> #1 < a ==> a < p - #1 ==> inv p a \<noteq> #1"
+ "p \<in> zprime ==> Numeral1 < a ==> a < p - Numeral1 ==> inv p a \<noteq> Numeral1"
-- {* same as @{text WilsonRuss} *}
apply safe
apply (cut_tac a = a and p = p in inv_is_inv)
prefer 4
apply simp
- apply (subgoal_tac "a = #1")
+ apply (subgoal_tac "a = Numeral1")
apply (rule_tac [2] zcong_zless_imp_eq)
apply auto
done
-lemma aux: "[a * (p - #1) = #1] (mod p) = [a = p - #1] (mod p)"
+lemma aux: "[a * (p - Numeral1) = Numeral1] (mod p) = [a = p - Numeral1] (mod p)"
-- {* same as @{text WilsonRuss} *}
apply (unfold zcong_def)
apply (simp add: zdiff_zdiff_eq zdiff_zdiff_eq2 zdiff_zmult_distrib2)
- apply (rule_tac s = "p dvd -((a + #1) + (p * -a))" in trans)
+ apply (rule_tac s = "p dvd -((a + Numeral1) + (p * -a))" in trans)
apply (simp add: zmult_commute zminus_zdiff_eq)
apply (subst zdvd_zminus_iff)
apply (subst zdvd_reduce)
- apply (rule_tac s = "p dvd (a + #1) + (p * -#1)" in trans)
+ apply (rule_tac s = "p dvd (a + Numeral1) + (p * -Numeral1)" in trans)
apply (subst zdvd_reduce)
apply auto
done
lemma inv_not_p_minus_1:
- "p \<in> zprime ==> #1 < a ==> a < p - #1 ==> inv p a \<noteq> p - #1"
+ "p \<in> zprime ==> Numeral1 < a ==> a < p - Numeral1 ==> inv p a \<noteq> p - Numeral1"
-- {* same as @{text WilsonRuss} *}
apply safe
apply (cut_tac a = a and p = p in inv_is_inv)
apply auto
apply (simp add: aux)
- apply (subgoal_tac "a = p - #1")
+ apply (subgoal_tac "a = p - Numeral1")
apply (rule_tac [2] zcong_zless_imp_eq)
apply auto
done
@@ -102,9 +102,9 @@
but use ``@{text correct}'' theorems.
*}
-lemma inv_g_1: "p \<in> zprime ==> #1 < a ==> a < p - #1 ==> #1 < inv p a"
- apply (subgoal_tac "inv p a \<noteq> #1")
- apply (subgoal_tac "inv p a \<noteq> #0")
+lemma inv_g_1: "p \<in> zprime ==> Numeral1 < a ==> a < p - Numeral1 ==> Numeral1 < inv p a"
+ apply (subgoal_tac "inv p a \<noteq> Numeral1")
+ apply (subgoal_tac "inv p a \<noteq> Numeral0")
apply (subst order_less_le)
apply (subst zle_add1_eq_le [symmetric])
apply (subst order_less_le)
@@ -116,7 +116,7 @@
done
lemma inv_less_p_minus_1:
- "p \<in> zprime ==> #1 < a ==> a < p - #1 ==> inv p a < p - #1"
+ "p \<in> zprime ==> Numeral1 < a ==> a < p - Numeral1 ==> inv p a < p - Numeral1"
-- {* ditto *}
apply (subst order_less_le)
apply (simp add: inv_not_p_minus_1 inv_less)
@@ -125,23 +125,23 @@
text {* \medskip Bijection *}
-lemma aux1: "#1 < x ==> #0 \<le> (x::int)"
+lemma aux1: "Numeral1 < x ==> Numeral0 \<le> (x::int)"
apply auto
done
-lemma aux2: "#1 < x ==> #0 < (x::int)"
+lemma aux2: "Numeral1 < x ==> Numeral0 < (x::int)"
apply auto
done
-lemma aux3: "x \<le> p - #2 ==> x < (p::int)"
+lemma aux3: "x \<le> p - # 2 ==> x < (p::int)"
apply auto
done
-lemma aux4: "x \<le> p - #2 ==> x < (p::int)-#1"
+lemma aux4: "x \<le> p - # 2 ==> x < (p::int)-Numeral1"
apply auto
done
-lemma inv_inj: "p \<in> zprime ==> inj_on (inv p) (d22set (p - #2))"
+lemma inv_inj: "p \<in> zprime ==> inj_on (inv p) (d22set (p - # 2))"
apply (unfold inj_on_def)
apply auto
apply (rule zcong_zless_imp_eq)
@@ -160,22 +160,22 @@
done
lemma inv_d22set_d22set:
- "p \<in> zprime ==> inv p ` d22set (p - #2) = d22set (p - #2)"
+ "p \<in> zprime ==> inv p ` d22set (p - # 2) = d22set (p - # 2)"
apply (rule endo_inj_surj)
apply (rule d22set_fin)
apply (erule_tac [2] inv_inj)
apply auto
apply (rule d22set_mem)
apply (erule inv_g_1)
- apply (subgoal_tac [3] "inv p xa < p - #1")
+ apply (subgoal_tac [3] "inv p xa < p - Numeral1")
apply (erule_tac [4] inv_less_p_minus_1)
apply (auto intro: d22set_g_1 d22set_le aux4)
done
lemma d22set_d22set_bij:
- "p \<in> zprime ==> (d22set (p - #2), d22set (p - #2)) \<in> bijR (reciR p)"
+ "p \<in> zprime ==> (d22set (p - # 2), d22set (p - # 2)) \<in> bijR (reciR p)"
apply (unfold reciR_def)
- apply (rule_tac s = "(d22set (p - #2), inv p ` d22set (p - #2))" in subst)
+ apply (rule_tac s = "(d22set (p - # 2), inv p ` d22set (p - # 2))" in subst)
apply (simp add: inv_d22set_d22set)
apply (rule inj_func_bijR)
apply (rule_tac [3] d22set_fin)
@@ -187,7 +187,7 @@
apply (auto intro: d22set_g_1 d22set_le aux2 aux3 aux4)
done
-lemma reciP_bijP: "p \<in> zprime ==> bijP (reciR p) (d22set (p - #2))"
+lemma reciP_bijP: "p \<in> zprime ==> bijP (reciR p) (d22set (p - # 2))"
apply (unfold reciR_def bijP_def)
apply auto
apply (rule d22set_mem)
@@ -217,7 +217,7 @@
apply auto
done
-lemma bijER_d22set: "p \<in> zprime ==> d22set (p - #2) \<in> bijER (reciR p)"
+lemma bijER_d22set: "p \<in> zprime ==> d22set (p - # 2) \<in> bijER (reciR p)"
apply (rule bijR_bijER)
apply (erule d22set_d22set_bij)
apply (erule reciP_bijP)
@@ -229,28 +229,28 @@
subsection {* Wilson *}
lemma bijER_zcong_prod_1:
- "p \<in> zprime ==> A \<in> bijER (reciR p) ==> [setprod A = #1] (mod p)"
+ "p \<in> zprime ==> A \<in> bijER (reciR p) ==> [setprod A = Numeral1] (mod p)"
apply (unfold reciR_def)
apply (erule bijER.induct)
- apply (subgoal_tac [2] "a = #1 \<or> a = p - #1")
+ apply (subgoal_tac [2] "a = Numeral1 \<or> a = p - Numeral1")
apply (rule_tac [3] zcong_square_zless)
apply auto
apply (subst setprod_insert)
prefer 3
apply (subst setprod_insert)
apply (auto simp add: fin_bijER)
- apply (subgoal_tac "zcong ((a * b) * setprod A) (#1 * #1) p")
+ apply (subgoal_tac "zcong ((a * b) * setprod A) (Numeral1 * Numeral1) p")
apply (simp add: zmult_assoc)
apply (rule zcong_zmult)
apply auto
done
-theorem Wilson_Bij: "p \<in> zprime ==> [zfact (p - #1) = #-1] (mod p)"
- apply (subgoal_tac "zcong ((p - #1) * zfact (p - #2)) (#-1 * #1) p")
+theorem Wilson_Bij: "p \<in> zprime ==> [zfact (p - Numeral1) = # -1] (mod p)"
+ apply (subgoal_tac "zcong ((p - Numeral1) * zfact (p - # 2)) (# -1 * Numeral1) p")
apply (rule_tac [2] zcong_zmult)
apply (simp add: zprime_def)
apply (subst zfact.simps)
- apply (rule_tac t = "p - #1 - #1" and s = "p - #2" in subst)
+ apply (rule_tac t = "p - Numeral1 - Numeral1" and s = "p - # 2" in subst)
apply auto
apply (simp add: zcong_def)
apply (subst d22set_prod_zfact [symmetric])
--- a/src/HOL/NumberTheory/WilsonRuss.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/NumberTheory/WilsonRuss.thy Fri Oct 05 21:52:39 2001 +0200
@@ -20,25 +20,25 @@
wset :: "int * int => int set"
defs
- inv_def: "inv p a == (a^(nat (p - #2))) mod p"
+ inv_def: "inv p a == (a^(nat (p - # 2))) mod p"
recdef wset
"measure ((\<lambda>(a, p). nat a) :: int * int => nat)"
"wset (a, p) =
- (if #1 < a then
- let ws = wset (a - #1, p)
+ (if Numeral1 < a then
+ let ws = wset (a - Numeral1, p)
in (if a \<in> ws then ws else insert a (insert (inv p a) ws)) else {})"
text {* \medskip @{term [source] inv} *}
-lemma aux: "#1 < m ==> Suc (nat (m - #2)) = nat (m - #1)"
+lemma aux: "Numeral1 < m ==> Suc (nat (m - # 2)) = nat (m - Numeral1)"
apply (subst int_int_eq [symmetric])
apply auto
done
lemma inv_is_inv:
- "p \<in> zprime \<Longrightarrow> #0 < a \<Longrightarrow> a < p ==> [a * inv p a = #1] (mod p)"
+ "p \<in> zprime \<Longrightarrow> Numeral0 < a \<Longrightarrow> a < p ==> [a * inv p a = Numeral1] (mod p)"
apply (unfold inv_def)
apply (subst zcong_zmod)
apply (subst zmod_zmult1_eq [symmetric])
@@ -52,71 +52,71 @@
done
lemma inv_distinct:
- "p \<in> zprime \<Longrightarrow> #1 < a \<Longrightarrow> a < p - #1 ==> a \<noteq> inv p a"
+ "p \<in> zprime \<Longrightarrow> Numeral1 < a \<Longrightarrow> a < p - Numeral1 ==> a \<noteq> inv p a"
apply safe
apply (cut_tac a = a and p = p in zcong_square)
apply (cut_tac [3] a = a and p = p in inv_is_inv)
apply auto
- apply (subgoal_tac "a = #1")
+ apply (subgoal_tac "a = Numeral1")
apply (rule_tac [2] m = p in zcong_zless_imp_eq)
- apply (subgoal_tac [7] "a = p - #1")
+ apply (subgoal_tac [7] "a = p - Numeral1")
apply (rule_tac [8] m = p in zcong_zless_imp_eq)
apply auto
done
lemma inv_not_0:
- "p \<in> zprime \<Longrightarrow> #1 < a \<Longrightarrow> a < p - #1 ==> inv p a \<noteq> #0"
+ "p \<in> zprime \<Longrightarrow> Numeral1 < a \<Longrightarrow> a < p - Numeral1 ==> inv p a \<noteq> Numeral0"
apply safe
apply (cut_tac a = a and p = p in inv_is_inv)
apply (unfold zcong_def)
apply auto
- apply (subgoal_tac "\<not> p dvd #1")
+ apply (subgoal_tac "\<not> p dvd Numeral1")
apply (rule_tac [2] zdvd_not_zless)
- apply (subgoal_tac "p dvd #1")
+ apply (subgoal_tac "p dvd Numeral1")
prefer 2
apply (subst zdvd_zminus_iff [symmetric])
apply auto
done
lemma inv_not_1:
- "p \<in> zprime \<Longrightarrow> #1 < a \<Longrightarrow> a < p - #1 ==> inv p a \<noteq> #1"
+ "p \<in> zprime \<Longrightarrow> Numeral1 < a \<Longrightarrow> a < p - Numeral1 ==> inv p a \<noteq> Numeral1"
apply safe
apply (cut_tac a = a and p = p in inv_is_inv)
prefer 4
apply simp
- apply (subgoal_tac "a = #1")
+ apply (subgoal_tac "a = Numeral1")
apply (rule_tac [2] zcong_zless_imp_eq)
apply auto
done
-lemma aux: "[a * (p - #1) = #1] (mod p) = [a = p - #1] (mod p)"
+lemma aux: "[a * (p - Numeral1) = Numeral1] (mod p) = [a = p - Numeral1] (mod p)"
apply (unfold zcong_def)
apply (simp add: zdiff_zdiff_eq zdiff_zdiff_eq2 zdiff_zmult_distrib2)
- apply (rule_tac s = "p dvd -((a + #1) + (p * -a))" in trans)
+ apply (rule_tac s = "p dvd -((a + Numeral1) + (p * -a))" in trans)
apply (simp add: zmult_commute zminus_zdiff_eq)
apply (subst zdvd_zminus_iff)
apply (subst zdvd_reduce)
- apply (rule_tac s = "p dvd (a + #1) + (p * -#1)" in trans)
+ apply (rule_tac s = "p dvd (a + Numeral1) + (p * -Numeral1)" in trans)
apply (subst zdvd_reduce)
apply auto
done
lemma inv_not_p_minus_1:
- "p \<in> zprime \<Longrightarrow> #1 < a \<Longrightarrow> a < p - #1 ==> inv p a \<noteq> p - #1"
+ "p \<in> zprime \<Longrightarrow> Numeral1 < a \<Longrightarrow> a < p - Numeral1 ==> inv p a \<noteq> p - Numeral1"
apply safe
apply (cut_tac a = a and p = p in inv_is_inv)
apply auto
apply (simp add: aux)
- apply (subgoal_tac "a = p - #1")
+ apply (subgoal_tac "a = p - Numeral1")
apply (rule_tac [2] zcong_zless_imp_eq)
apply auto
done
lemma inv_g_1:
- "p \<in> zprime \<Longrightarrow> #1 < a \<Longrightarrow> a < p - #1 ==> #1 < inv p a"
- apply (case_tac "#0\<le> inv p a")
- apply (subgoal_tac "inv p a \<noteq> #1")
- apply (subgoal_tac "inv p a \<noteq> #0")
+ "p \<in> zprime \<Longrightarrow> Numeral1 < a \<Longrightarrow> a < p - Numeral1 ==> Numeral1 < inv p a"
+ apply (case_tac "Numeral0\<le> inv p a")
+ apply (subgoal_tac "inv p a \<noteq> Numeral1")
+ apply (subgoal_tac "inv p a \<noteq> Numeral0")
apply (subst order_less_le)
apply (subst zle_add1_eq_le [symmetric])
apply (subst order_less_le)
@@ -128,7 +128,7 @@
done
lemma inv_less_p_minus_1:
- "p \<in> zprime \<Longrightarrow> #1 < a \<Longrightarrow> a < p - #1 ==> inv p a < p - #1"
+ "p \<in> zprime \<Longrightarrow> Numeral1 < a \<Longrightarrow> a < p - Numeral1 ==> inv p a < p - Numeral1"
apply (case_tac "inv p a < p")
apply (subst order_less_le)
apply (simp add: inv_not_p_minus_1)
@@ -137,24 +137,24 @@
apply (simp add: pos_mod_bound)
done
-lemma aux: "#5 \<le> p ==>
- nat (p - #2) * nat (p - #2) = Suc (nat (p - #1) * nat (p - #3))"
+lemma aux: "# 5 \<le> p ==>
+ nat (p - # 2) * nat (p - # 2) = Suc (nat (p - Numeral1) * nat (p - # 3))"
apply (subst int_int_eq [symmetric])
apply (simp add: zmult_int [symmetric])
apply (simp add: zdiff_zmult_distrib zdiff_zmult_distrib2)
done
lemma zcong_zpower_zmult:
- "[x^y = #1] (mod p) \<Longrightarrow> [x^(y * z) = #1] (mod p)"
+ "[x^y = Numeral1] (mod p) \<Longrightarrow> [x^(y * z) = Numeral1] (mod p)"
apply (induct z)
apply (auto simp add: zpower_zadd_distrib)
- apply (subgoal_tac "zcong (x^y * x^(y * n)) (#1 * #1) p")
+ apply (subgoal_tac "zcong (x^y * x^(y * n)) (Numeral1 * Numeral1) p")
apply (rule_tac [2] zcong_zmult)
apply simp_all
done
lemma inv_inv: "p \<in> zprime \<Longrightarrow>
- #5 \<le> p \<Longrightarrow> #0 < a \<Longrightarrow> a < p ==> inv p (inv p a) = a"
+ # 5 \<le> p \<Longrightarrow> Numeral0 < a \<Longrightarrow> a < p ==> inv p (inv p a) = a"
apply (unfold inv_def)
apply (subst zpower_zmod)
apply (subst zpower_zpower)
@@ -165,7 +165,7 @@
apply (subst zcong_zmod [symmetric])
apply (subst aux)
apply (subgoal_tac [2]
- "zcong (a * a^(nat (p - #1) * nat (p - #3))) (a * #1) p")
+ "zcong (a * a^(nat (p - Numeral1) * nat (p - # 3))) (a * Numeral1) p")
apply (rule_tac [3] zcong_zmult)
apply (rule_tac [4] zcong_zpower_zmult)
apply (erule_tac [4] Little_Fermat)
@@ -180,7 +180,7 @@
lemma wset_induct:
"(!!a p. P {} a p) \<Longrightarrow>
- (!!a p. #1 < (a::int) \<Longrightarrow> P (wset (a - #1, p)) (a - #1) p
+ (!!a p. Numeral1 < (a::int) \<Longrightarrow> P (wset (a - Numeral1, p)) (a - Numeral1) p
==> P (wset (a, p)) a p)
==> P (wset (u, v)) u v"
proof -
@@ -188,7 +188,7 @@
show ?thesis
apply (rule wset.induct)
apply safe
- apply (case_tac [2] "#1 < a")
+ apply (case_tac [2] "Numeral1 < a")
apply (rule_tac [2] rule_context)
apply simp_all
apply (simp_all add: wset.simps rule_context)
@@ -196,27 +196,27 @@
qed
lemma wset_mem_imp_or [rule_format]:
- "#1 < a \<Longrightarrow> b \<notin> wset (a - #1, p)
+ "Numeral1 < a \<Longrightarrow> b \<notin> wset (a - Numeral1, p)
==> b \<in> wset (a, p) --> b = a \<or> b = inv p a"
apply (subst wset.simps)
apply (unfold Let_def)
apply simp
done
-lemma wset_mem_mem [simp]: "#1 < a ==> a \<in> wset (a, p)"
+lemma wset_mem_mem [simp]: "Numeral1 < a ==> a \<in> wset (a, p)"
apply (subst wset.simps)
apply (unfold Let_def)
apply simp
done
-lemma wset_subset: "#1 < a \<Longrightarrow> b \<in> wset (a - #1, p) ==> b \<in> wset (a, p)"
+lemma wset_subset: "Numeral1 < a \<Longrightarrow> b \<in> wset (a - Numeral1, p) ==> b \<in> wset (a, p)"
apply (subst wset.simps)
apply (unfold Let_def)
apply auto
done
lemma wset_g_1 [rule_format]:
- "p \<in> zprime --> a < p - #1 --> b \<in> wset (a, p) --> #1 < b"
+ "p \<in> zprime --> a < p - Numeral1 --> b \<in> wset (a, p) --> Numeral1 < b"
apply (induct a p rule: wset_induct)
apply auto
apply (case_tac "b = a")
@@ -230,7 +230,7 @@
done
lemma wset_less [rule_format]:
- "p \<in> zprime --> a < p - #1 --> b \<in> wset (a, p) --> b < p - #1"
+ "p \<in> zprime --> a < p - Numeral1 --> b \<in> wset (a, p) --> b < p - Numeral1"
apply (induct a p rule: wset_induct)
apply auto
apply (case_tac "b = a")
@@ -245,7 +245,7 @@
lemma wset_mem [rule_format]:
"p \<in> zprime -->
- a < p - #1 --> #1 < b --> b \<le> a --> b \<in> wset (a, p)"
+ a < p - Numeral1 --> Numeral1 < b --> b \<le> a --> b \<in> wset (a, p)"
apply (induct a p rule: wset.induct)
apply auto
apply (subgoal_tac "b = a")
@@ -256,7 +256,7 @@
done
lemma wset_mem_inv_mem [rule_format]:
- "p \<in> zprime --> #5 \<le> p --> a < p - #1 --> b \<in> wset (a, p)
+ "p \<in> zprime --> # 5 \<le> p --> a < p - Numeral1 --> b \<in> wset (a, p)
--> inv p b \<in> wset (a, p)"
apply (induct a p rule: wset_induct)
apply auto
@@ -274,7 +274,7 @@
done
lemma wset_inv_mem_mem:
- "p \<in> zprime \<Longrightarrow> #5 \<le> p \<Longrightarrow> a < p - #1 \<Longrightarrow> #1 < b \<Longrightarrow> b < p - #1
+ "p \<in> zprime \<Longrightarrow> # 5 \<le> p \<Longrightarrow> a < p - Numeral1 \<Longrightarrow> Numeral1 < b \<Longrightarrow> b < p - Numeral1
\<Longrightarrow> inv p b \<in> wset (a, p) \<Longrightarrow> b \<in> wset (a, p)"
apply (rule_tac s = "inv p (inv p b)" and t = b in subst)
apply (rule_tac [2] wset_mem_inv_mem)
@@ -292,7 +292,7 @@
lemma wset_zcong_prod_1 [rule_format]:
"p \<in> zprime -->
- #5 \<le> p --> a < p - #1 --> [setprod (wset (a, p)) = #1] (mod p)"
+ # 5 \<le> p --> a < p - Numeral1 --> [setprod (wset (a, p)) = Numeral1] (mod p)"
apply (induct a p rule: wset_induct)
prefer 2
apply (subst wset.simps)
@@ -301,20 +301,20 @@
apply (subst setprod_insert)
apply (tactic {* stac (thm "setprod_insert") 3 *})
apply (subgoal_tac [5]
- "zcong (a * inv p a * setprod (wset (a - #1, p))) (#1 * #1) p")
+ "zcong (a * inv p a * setprod (wset (a - Numeral1, p))) (Numeral1 * Numeral1) p")
prefer 5
apply (simp add: zmult_assoc)
apply (rule_tac [5] zcong_zmult)
apply (rule_tac [5] inv_is_inv)
apply (tactic "Clarify_tac 4")
- apply (subgoal_tac [4] "a \<in> wset (a - #1, p)")
+ apply (subgoal_tac [4] "a \<in> wset (a - Numeral1, p)")
apply (rule_tac [5] wset_inv_mem_mem)
apply (simp_all add: wset_fin)
apply (rule inv_distinct)
apply auto
done
-lemma d22set_eq_wset: "p \<in> zprime ==> d22set (p - #2) = wset (p - #2, p)"
+lemma d22set_eq_wset: "p \<in> zprime ==> d22set (p - # 2) = wset (p - # 2, p)"
apply safe
apply (erule wset_mem)
apply (rule_tac [2] d22set_g_1)
@@ -323,7 +323,7 @@
apply (erule_tac [4] wset_g_1)
prefer 6
apply (subst zle_add1_eq_le [symmetric])
- apply (subgoal_tac "p - #2 + #1 = p - #1")
+ apply (subgoal_tac "p - # 2 + Numeral1 = p - Numeral1")
apply (simp (no_asm_simp))
apply (erule wset_less)
apply auto
@@ -332,36 +332,36 @@
subsection {* Wilson *}
-lemma prime_g_5: "p \<in> zprime \<Longrightarrow> p \<noteq> #2 \<Longrightarrow> p \<noteq> #3 ==> #5 \<le> p"
+lemma prime_g_5: "p \<in> zprime \<Longrightarrow> p \<noteq> # 2 \<Longrightarrow> p \<noteq> # 3 ==> # 5 \<le> p"
apply (unfold zprime_def dvd_def)
- apply (case_tac "p = #4")
+ apply (case_tac "p = # 4")
apply auto
apply (rule notE)
prefer 2
apply assumption
apply (simp (no_asm))
- apply (rule_tac x = "#2" in exI)
+ apply (rule_tac x = "# 2" in exI)
apply safe
- apply (rule_tac x = "#2" in exI)
+ apply (rule_tac x = "# 2" in exI)
apply auto
apply arith
done
theorem Wilson_Russ:
- "p \<in> zprime ==> [zfact (p - #1) = #-1] (mod p)"
- apply (subgoal_tac "[(p - #1) * zfact (p - #2) = #-1 * #1] (mod p)")
+ "p \<in> zprime ==> [zfact (p - Numeral1) = # -1] (mod p)"
+ apply (subgoal_tac "[(p - Numeral1) * zfact (p - # 2) = # -1 * Numeral1] (mod p)")
apply (rule_tac [2] zcong_zmult)
apply (simp only: zprime_def)
apply (subst zfact.simps)
- apply (rule_tac t = "p - #1 - #1" and s = "p - #2" in subst)
+ apply (rule_tac t = "p - Numeral1 - Numeral1" and s = "p - # 2" in subst)
apply auto
apply (simp only: zcong_def)
apply (simp (no_asm_simp))
- apply (case_tac "p = #2")
+ apply (case_tac "p = # 2")
apply (simp add: zfact.simps)
- apply (case_tac "p = #3")
+ apply (case_tac "p = # 3")
apply (simp add: zfact.simps)
- apply (subgoal_tac "#5 \<le> p")
+ apply (subgoal_tac "# 5 \<le> p")
apply (erule_tac [2] prime_g_5)
apply (subst d22set_prod_zfact [symmetric])
apply (subst d22set_eq_wset)
--- a/src/HOL/Power.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Power.ML Fri Oct 05 21:52:39 2001 +0200
@@ -30,7 +30,7 @@
by Auto_tac;
qed "power_eq_0D";
-Goal "!!i::nat. 1 <= i ==> 1' <= i^n";
+Goal "!!i::nat. 1 <= i ==> Suc 0 <= i^n";
by (induct_tac "n" 1);
by Auto_tac;
qed "one_le_power";
@@ -49,12 +49,12 @@
by (asm_simp_tac (simpset() addsimps [power_add]) 1);
qed "le_imp_power_dvd";
-Goal "1 < i ==> \\<forall>n. i ^ m <= i ^ n --> m <= n";
+Goal "(1::nat) < i ==> \\<forall>n. i ^ m <= i ^ n --> m <= n";
by (induct_tac "m" 1);
by Auto_tac;
by (case_tac "na" 1);
by Auto_tac;
-by (subgoal_tac "2 * 1 <= i * i^n" 1);
+by (subgoal_tac "Suc 1 * 1 <= i * i^n" 1);
by (Asm_full_simp_tac 1);
by (rtac mult_le_mono 1);
by Auto_tac;
@@ -73,7 +73,7 @@
by (blast_tac (claset() addSDs [dvd_mult_right]) 1);
qed_spec_mp "power_le_dvd";
-Goal "[|i^m dvd i^n; 1 < i|] ==> m <= n";
+Goal "[|i^m dvd i^n; (1::nat) < i|] ==> m <= n";
by (rtac power_le_imp_le 1);
by (assume_tac 1);
by (etac dvd_imp_le 1);
@@ -120,7 +120,7 @@
qed "binomial_Suc_n";
Addsimps [binomial_Suc_n];
-Goal "(n choose 1') = n";
+Goal "(n choose Suc 0) = n";
by (induct_tac "n" 1);
by (ALLGOALS Asm_simp_tac);
qed "binomial_1";
@@ -162,8 +162,8 @@
qed "binomial_Suc_Suc_eq_times";
(*Another version, with -1 instead of Suc.*)
-Goal "[|k <= n; 0<k|] ==> (n choose k) * k = n * ((n-1) choose (k-1))";
-by (cut_inst_tac [("n","n-1"),("k","k-1")] Suc_times_binomial_eq 1);
+Goal "[|k <= n; 0<k|] ==> (n choose k) * k = n * ((n - 1) choose (k - 1))";
+by (cut_inst_tac [("n","n - 1"),("k","k - 1")] Suc_times_binomial_eq 1);
by (asm_full_simp_tac (simpset() addsplits [nat_diff_split]) 1);
by Auto_tac;
qed "times_binomial_minus1_eq";
--- a/src/HOL/Prolog/Type.ML Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Prolog/Type.ML Fri Oct 05 21:52:39 2001 +0200
@@ -6,7 +6,7 @@
pgoal "typeof (fix (%x. x)) ?T";
-pgoal "typeof (fix (%fact. abs(%n. (app fact (n-0))))) ?T";
+pgoal "typeof (fix (%fact. abs(%n. (app fact (n - 0))))) ?T";
pgoal "typeof (fix (%fact. abs(%n. cond (n eq 0) (S 0) \
\(n * (app fact (n - (S 0))))))) ?T";
--- a/src/HOL/Real/HahnBanach/Aux.thy Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/Real/HahnBanach/Aux.thy Fri Oct 05 21:52:39 2001 +0200
@@ -40,20 +40,20 @@
text {* \medskip Some lemmas for the reals. *}
-lemma real_add_minus_eq: "x - y = (#0::real) \<Longrightarrow> x = y"
+lemma real_add_minus_eq: "x - y = (Numeral0::real) \<Longrightarrow> x = y"
by simp
-lemma abs_minus_one: "abs (- (#1::real)) = #1"
+lemma abs_minus_one: "abs (- (Numeral1::real)) = Numeral1"
by simp
lemma real_mult_le_le_mono1a:
- "(#0::real) \<le> z \<Longrightarrow> x \<le> y \<Longrightarrow> z * x \<le> z * y"
+ "(Numeral0::real) \<le> z \<Longrightarrow> x \<le> y \<Longrightarrow> z * x \<le> z * y"
by (simp add: real_mult_le_mono2)
lemma real_mult_le_le_mono2:
- "(#0::real) \<le> z \<Longrightarrow> x \<le> y \<Longrightarrow> x * z \<le> y * z"
+ "(Numeral0::real) \<le> z \<Longrightarrow> x \<le> y \<Longrightarrow> x * z \<le> y * z"
proof -
- assume "(#0::real) \<le> z" "x \<le> y"
+ assume "(Numeral0::real) \<le> z" "x \<le> y"
hence "x < y \<or> x = y" by (auto simp add: order_le_less)
thus ?thesis
proof
@@ -66,11 +66,11 @@
qed
lemma real_mult_less_le_anti:
- "z < (#0::real) \<Longrightarrow> x \<le> y \<Longrightarrow> z * y \<le> z * x"
+ "z < (Numeral0::real) \<Longrightarrow> x \<le> y \<Longrightarrow> z * y \<le> z * x"
proof -