author | wenzelm |
Fri, 06 Oct 2000 17:35:58 +0200 | |
changeset 10168 | 50be659d4222 |
parent 10045 | c76b73e16711 |
child 10607 | 352f6f209775 |
permissions | -rw-r--r-- |
10045
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
1 |
(* Title : HRealAbs.ML |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
2 |
Author : Jacques D. Fleuriot |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
3 |
Copyright : 1998 University of Cambridge |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
4 |
Description : Absolute value function for the hyperreals |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
5 |
Similar to RealAbs.thy |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
6 |
*) |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
7 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
8 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
9 |
(*------------------------------------------------------------ |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
10 |
absolute value on hyperreals as pointwise operation on |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
11 |
equivalence class representative |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
12 |
------------------------------------------------------------*) |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
13 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
14 |
Goalw [hrabs_def] |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
15 |
"abs (Abs_hypreal (hyprel ^^ {X})) = \ |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
16 |
\ Abs_hypreal(hyprel ^^ {%n. abs (X n)})"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
17 |
by (auto_tac (claset(),simpset() addsimps [hypreal_zero_def, |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
18 |
hypreal_le,hypreal_minus])); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
19 |
by (ALLGOALS(Ultra_tac THEN' arith_tac )); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
20 |
qed "hypreal_hrabs"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
21 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
22 |
(*------------------------------------------------------------ |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
23 |
Properties of the absolute value function over the reals |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
24 |
(adapted version of previously proved theorems about abs) |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
25 |
------------------------------------------------------------*) |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
26 |
Goalw [hrabs_def] "abs r = (if (0::hypreal)<=r then r else -r)"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
27 |
by (Step_tac 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
28 |
qed "hrabs_iff"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
29 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
30 |
Goalw [hrabs_def] "abs (0::hypreal) = (0::hypreal)"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
31 |
by (rtac (hypreal_le_refl RS if_P) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
32 |
qed "hrabs_zero"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
33 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
34 |
Addsimps [hrabs_zero]; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
35 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
36 |
Goalw [hrabs_def] "abs (0::hypreal) = -(0::hypreal)"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
37 |
by (rtac (hypreal_minus_zero RS ssubst) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
38 |
by (rtac if_cancel 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
39 |
qed "hrabs_minus_zero"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
40 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
41 |
val [prem] = goalw thy [hrabs_def] "(0::hypreal)<=x ==> abs x = x"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
42 |
by (rtac (prem RS if_P) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
43 |
qed "hrabs_eqI1"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
44 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
45 |
val [prem] = goalw thy [hrabs_def] "(0::hypreal)<x ==> abs x = x"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
46 |
by (simp_tac (simpset() addsimps [(prem |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
47 |
RS hypreal_less_imp_le),hrabs_eqI1]) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
48 |
qed "hrabs_eqI2"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
49 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
50 |
val [prem] = goalw thy [hrabs_def,hypreal_le_def] |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
51 |
"x<(0::hypreal) ==> abs x = -x"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
52 |
by (simp_tac (simpset() addsimps [prem,if_not_P]) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
53 |
qed "hrabs_minus_eqI2"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
54 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
55 |
Goal "!!x. x<=(0::hypreal) ==> abs x = -x"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
56 |
by (dtac hypreal_le_imp_less_or_eq 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
57 |
by (fast_tac (HOL_cs addIs [hrabs_minus_zero, |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
58 |
hrabs_minus_eqI2]) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
59 |
qed "hrabs_minus_eqI1"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
60 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
61 |
Goalw [hrabs_def,hypreal_le_def] "(0::hypreal)<= abs x"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
62 |
by (auto_tac (claset() addDs [hypreal_minus_zero_less_iff RS iffD2, |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
63 |
hypreal_less_asym],simpset())); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
64 |
qed "hrabs_ge_zero"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
65 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
66 |
Goal "abs(abs x)=abs (x::hypreal)"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
67 |
by (res_inst_tac [("r1","abs x")] (hrabs_iff RS ssubst) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
68 |
by (blast_tac (claset() addIs [if_P,hrabs_ge_zero]) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
69 |
qed "hrabs_idempotent"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
70 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
71 |
Goalw [hrabs_def] "(x=(0::hypreal)) = (abs x = (0::hypreal))"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
72 |
by (Simp_tac 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
73 |
qed "hrabs_zero_iff"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
74 |
Addsimps [hrabs_zero_iff RS sym]; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
75 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
76 |
Goal "(x ~= (0::hypreal)) = (abs x ~= 0)"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
77 |
by (Simp_tac 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
78 |
qed "hrabs_not_zero_iff"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
79 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
80 |
Goalw [hrabs_def] "(x::hypreal)<=abs x"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
81 |
by (auto_tac (claset() addDs [not_hypreal_leE RS hypreal_less_imp_le], |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
82 |
simpset() addsimps [hypreal_le_zero_iff])); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
83 |
qed "hrabs_ge_self"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
84 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
85 |
Goalw [hrabs_def] "-(x::hypreal)<=abs x"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
86 |
by (full_simp_tac (simpset() addsimps [hypreal_ge_zero_iff]) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
87 |
qed "hrabs_ge_minus_self"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
88 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
89 |
(* very short proof by "transfer" *) |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
90 |
Goal "abs(x*(y::hypreal)) = (abs x)*(abs y)"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
91 |
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
92 |
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
93 |
by (auto_tac (claset(),simpset() addsimps [hypreal_hrabs, |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
94 |
hypreal_mult,abs_mult])); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
95 |
qed "hrabs_mult"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
96 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
97 |
Goal "x~= (0::hypreal) ==> abs(hrinv(x)) = hrinv(abs(x))"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
98 |
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
99 |
by (auto_tac (claset(),simpset() addsimps [hypreal_hrabs, |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
100 |
hypreal_hrinv,hypreal_zero_def])); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
101 |
by (ultra_tac (claset(),simpset() addsimps [abs_rinv]) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
102 |
by (arith_tac 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
103 |
qed "hrabs_hrinv"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
104 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
105 |
(* old version of proof: |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
106 |
Goalw [hrabs_def] |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
107 |
"x~= (0::hypreal) ==> abs(hrinv(x)) = hrinv(abs(x))"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
108 |
by (auto_tac (claset(),simpset() addsimps [hypreal_minus_hrinv])); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
109 |
by (ALLGOALS(dtac not_hypreal_leE)); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
110 |
by (etac hypreal_less_asym 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
111 |
by (blast_tac (claset() addDs [hypreal_le_imp_less_or_eq, |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
112 |
hypreal_hrinv_gt_zero]) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
113 |
by (dtac (hrinv_not_zero RS not_sym) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
114 |
by (rtac (hypreal_hrinv_less_zero RSN (2,hypreal_less_asym)) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
115 |
by (assume_tac 2); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
116 |
by (blast_tac (claset() addSDs [hypreal_le_imp_less_or_eq]) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
117 |
qed "hrabs_hrinv"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
118 |
*) |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
119 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
120 |
val [prem] = goal thy "y ~= (0::hypreal) ==> \ |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
121 |
\ abs(x*hrinv(y)) = abs(x)*hrinv(abs(y))"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
122 |
by (res_inst_tac [("c1","abs y")] (hypreal_mult_left_cancel RS subst) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
123 |
by (simp_tac (simpset() addsimps [(hrabs_not_zero_iff RS sym), prem]) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
124 |
by (simp_tac (simpset() addsimps [(hrabs_mult RS sym), prem, |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
125 |
hrabs_not_zero_iff RS sym] @ hypreal_mult_ac) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
126 |
qed "hrabs_mult_hrinv"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
127 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
128 |
Goal "abs(x+(y::hypreal)) <= abs x + abs y"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
129 |
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
130 |
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
131 |
by (auto_tac (claset(),simpset() addsimps [hypreal_hrabs, |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
132 |
hypreal_add,hypreal_le,abs_triangle_ineq])); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
133 |
qed "hrabs_triangle_ineq"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
134 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
135 |
Goal "abs((w::hypreal) + x + y) <= abs(w) + abs(x) + abs(y)"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
136 |
by (auto_tac (claset() addSIs [(hrabs_triangle_ineq |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
137 |
RS hypreal_le_trans),hypreal_add_left_le_mono1], |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
138 |
simpset() addsimps [hypreal_add_assoc])); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
139 |
qed "hrabs_triangle_ineq_three"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
140 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
141 |
Goalw [hrabs_def] "abs(-x)=abs((x::hypreal))"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
142 |
by (auto_tac (claset() addSDs [not_hypreal_leE, |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
143 |
hypreal_less_asym] addIs [hypreal_le_anti_sym], |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
144 |
simpset() addsimps [hypreal_ge_zero_iff])); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
145 |
qed "hrabs_minus_cancel"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
146 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
147 |
Goal "abs((x::hypreal) + -y) = abs (y + -x)"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
148 |
by (rtac (hrabs_minus_cancel RS subst) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
149 |
by (simp_tac (simpset() addsimps [hypreal_add_commute]) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
150 |
qed "hrabs_minus_add_cancel"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
151 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
152 |
Goal "abs((x::hypreal) + -y) <= abs x + abs y"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
153 |
by (res_inst_tac [("x1","y")] (hrabs_minus_cancel RS subst) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
154 |
by (rtac hrabs_triangle_ineq 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
155 |
qed "rhabs_triangle_minus_ineq"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
156 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
157 |
val prem1::prem2::rest = goal thy |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
158 |
"[| abs x < r; abs y < s |] ==> abs(x+y) < r + (s::hypreal)"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
159 |
by (rtac hypreal_le_less_trans 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
160 |
by (rtac hrabs_triangle_ineq 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
161 |
by (rtac ([prem1,prem2] MRS hypreal_add_less_mono) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
162 |
qed "hrabs_add_less"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
163 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
164 |
Goal "[| abs x < r; abs y < s |] \ |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
165 |
\ ==> abs(x+ -y) < r + (s::hypreal)"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
166 |
by (rotate_tac 1 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
167 |
by (dtac (hrabs_minus_cancel RS ssubst) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
168 |
by (asm_simp_tac (simpset() addsimps [hrabs_add_less]) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
169 |
qed "hrabs_add_minus_less"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
170 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
171 |
val prem1::prem2::rest = |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
172 |
goal thy "[| abs x<r; abs y<s |] ==> abs(x*y)<r*(s::hypreal)"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
173 |
by (simp_tac (simpset() addsimps [hrabs_mult]) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
174 |
by (rtac hypreal_mult_le_less_trans 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
175 |
by (rtac hrabs_ge_zero 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
176 |
by (rtac prem2 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
177 |
by (rtac hypreal_mult_less_mono1 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
178 |
by (rtac (prem2 RS (hrabs_ge_zero RS hypreal_le_less_trans)) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
179 |
by (rtac prem1 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
180 |
by (rtac ([prem1 RS (hrabs_ge_zero RS hypreal_le_less_trans), |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
181 |
prem2 RS (hrabs_ge_zero RS hypreal_le_less_trans)] |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
182 |
MRS hypreal_mult_order) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
183 |
qed "hrabs_mult_less"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
184 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
185 |
Goal "!! x y r. 1hr < abs x ==> abs y <= abs(x*y)"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
186 |
by (cut_inst_tac [("x1","y")] (hrabs_ge_zero RS hypreal_le_imp_less_or_eq) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
187 |
by (EVERY1[etac disjE,rtac hypreal_less_imp_le]); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
188 |
by (dres_inst_tac [("x1","1hr")] (hypreal_less_minus_iff RS iffD1) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
189 |
by (forw_inst_tac [("y","abs x +-1hr")] hypreal_mult_order 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
190 |
by (assume_tac 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
191 |
by (rtac (hypreal_less_minus_iff RS iffD2) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
192 |
by (asm_full_simp_tac (simpset() addsimps [hypreal_add_mult_distrib2, |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
193 |
hrabs_mult, hypreal_mult_commute,hypreal_minus_mult_eq2 RS sym]) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
194 |
by (dtac sym 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
195 |
by (asm_full_simp_tac (simpset() addsimps [hypreal_le_refl,hrabs_mult]) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
196 |
qed "hrabs_mult_le"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
197 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
198 |
Goal "!!x. [| 1hr < abs x; r < abs y|] ==> r < abs(x*y)"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
199 |
by (fast_tac (HOL_cs addIs [hrabs_mult_le, hypreal_less_le_trans]) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
200 |
qed "hrabs_mult_gt"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
201 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
202 |
Goal "!!r. abs x < r ==> (0::hypreal) < r"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
203 |
by (blast_tac (claset() addSIs [hypreal_le_less_trans, |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
204 |
hrabs_ge_zero]) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
205 |
qed "hrabs_less_gt_zero"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
206 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
207 |
Goalw [hrabs_def] "abs 1hr = 1hr"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
208 |
by (auto_tac (claset() addSDs [not_hypreal_leE |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
209 |
RS hypreal_less_asym],simpset() addsimps |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
210 |
[hypreal_zero_less_one])); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
211 |
qed "hrabs_one"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
212 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
213 |
val prem1::prem2::rest = |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
214 |
goal thy "[| (0::hypreal) < x ; x < r |] ==> abs x < r"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
215 |
by (simp_tac (simpset() addsimps [(prem1 RS hrabs_eqI2),prem2]) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
216 |
qed "hrabs_lessI"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
217 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
218 |
Goal "abs x = (x::hypreal) | abs x = -x"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
219 |
by (cut_inst_tac [("x","0"),("y","x")] hypreal_linear 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
220 |
by (fast_tac (claset() addIs [hrabs_eqI2,hrabs_minus_eqI2, |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
221 |
hrabs_zero,hrabs_minus_zero]) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
222 |
qed "hrabs_disj"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
223 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
224 |
Goal "abs x = (y::hypreal) ==> x = y | -x = y"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
225 |
by (dtac sym 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
226 |
by (hyp_subst_tac 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
227 |
by (res_inst_tac [("x1","x")] (hrabs_disj RS disjE) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
228 |
by (REPEAT(Asm_simp_tac 1)); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
229 |
qed "hrabs_eq_disj"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
230 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
231 |
Goal "(abs x < (r::hypreal)) = (-r < x & x < r)"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
232 |
by (Step_tac 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
233 |
by (rtac (hypreal_less_swap_iff RS iffD2) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
234 |
by (asm_simp_tac (simpset() addsimps [(hrabs_ge_minus_self |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
235 |
RS hypreal_le_less_trans)]) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
236 |
by (asm_simp_tac (simpset() addsimps [(hrabs_ge_self |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
237 |
RS hypreal_le_less_trans)]) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
238 |
by (EVERY1 [dtac (hypreal_less_swap_iff RS iffD1), rotate_tac 1, |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
239 |
dtac (hypreal_minus_minus RS subst), |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
240 |
cut_inst_tac [("x","x")] hrabs_disj, dtac disjE ]); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
241 |
by (assume_tac 3 THEN Auto_tac); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
242 |
qed "hrabs_interval_iff"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
243 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
244 |
Goal "(abs x < (r::hypreal)) = (- x < r & x < r)"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
245 |
by (auto_tac (claset(),simpset() addsimps [hrabs_interval_iff])); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
246 |
by (dtac (hypreal_less_swap_iff RS iffD1) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
247 |
by (dtac (hypreal_less_swap_iff RS iffD1) 2); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
248 |
by (Auto_tac); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
249 |
qed "hrabs_interval_iff2"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
250 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
251 |
Goal |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
252 |
"(abs (x + -y) < (r::hypreal)) = (y + -r < x & x < y + r)"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
253 |
by (auto_tac (claset(),simpset() addsimps |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
254 |
[hrabs_interval_iff])); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
255 |
by (ALLGOALS(dtac (hypreal_less_minus_iff RS iffD1))); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
256 |
by (ALLGOALS(dtac (hypreal_less_minus_iff RS iffD1))); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
257 |
by (ALLGOALS(rtac (hypreal_less_minus_iff RS iffD2))); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
258 |
by (auto_tac (claset(),simpset() addsimps |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
259 |
[hypreal_minus_add_distrib] addsimps hypreal_add_ac)); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
260 |
qed "hrabs_add_minus_interval_iff"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
261 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
262 |
Goal "x < (y::hypreal) ==> abs(y + -x) = y + -x"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
263 |
by (dtac (hypreal_less_minus_iff RS iffD1) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
264 |
by (etac hrabs_eqI2 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
265 |
qed "hrabs_less_eqI2"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
266 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
267 |
Goal "x < (y::hypreal) ==> abs(x + -y) = y + -x"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
268 |
by (auto_tac (claset() addDs [hrabs_less_eqI2], |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
269 |
simpset() addsimps [hrabs_minus_add_cancel])); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
270 |
qed "hrabs_less_eqI2a"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
271 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
272 |
Goal "x <= (y::hypreal) ==> abs(y + -x) = y + -x"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
273 |
by (auto_tac (claset() addDs [hypreal_le_imp_less_or_eq, |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
274 |
hrabs_less_eqI2],simpset())); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
275 |
qed "hrabs_le_eqI2"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
276 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
277 |
Goal "x <= (y::hypreal) ==> abs(x + -y) = y + -x"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
278 |
by (auto_tac (claset() addDs [hrabs_le_eqI2], |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
279 |
simpset() addsimps [hrabs_minus_add_cancel])); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
280 |
qed "hrabs_le_eqI2a"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
281 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
282 |
(* Needed in Geom.ML *) |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
283 |
Goal "(y::hypreal) + - x + (y + - z) = abs (x + - z) \ |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
284 |
\ ==> y = z | x = y"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
285 |
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
286 |
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
287 |
by (res_inst_tac [("z","z")] eq_Abs_hypreal 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
288 |
by (auto_tac (claset(),simpset() addsimps [hypreal_hrabs, |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
289 |
hypreal_minus,hypreal_add])); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
290 |
by (Ultra_tac 1 THEN arith_tac 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
291 |
qed "hrabs_add_lemma_disj"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
292 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
293 |
(* Needed in Geom.ML *) |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
294 |
Goal "(x::hypreal) + - y + (z + - y) = abs (x + - z) \ |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
295 |
\ ==> y = z | x = y"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
296 |
by (rtac (hypreal_minus_eq_cancel RS subst) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
297 |
by (res_inst_tac [("b1","y")] (hypreal_minus_eq_cancel RS subst) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
298 |
by (rtac hrabs_add_lemma_disj 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
299 |
by (asm_full_simp_tac (simpset() addsimps [hrabs_minus_add_cancel] |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
300 |
@ hypreal_add_ac) 1); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
301 |
qed "hrabs_add_lemma_disj2"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
302 |
|
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
303 |
(*---------------------------------------------------------- |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
304 |
Relating hrabs to abs through embedding of IR into IR* |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
305 |
----------------------------------------------------------*) |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
306 |
Goalw [hypreal_of_real_def] |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
307 |
"abs (hypreal_of_real r) = hypreal_of_real (abs r)"; |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
308 |
by (auto_tac (claset(),simpset() addsimps [hypreal_hrabs])); |
c76b73e16711
New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff
changeset
|
309 |
qed "hypreal_of_real_hrabs"; |