src/HOL/Real/Hyperreal/HyperNat.ML
changeset 10751 a81ea5d3dd41
parent 10750 a681d3df1a39
child 10752 c4f1bf2acf4c
equal deleted inserted replaced
10750:a681d3df1a39 10751:a81ea5d3dd41
     1 (*  Title       : HyperNat.ML
       
     2     Author      : Jacques D. Fleuriot
       
     3     Copyright   : 1998  University of Cambridge
       
     4     Description : Explicit construction of hypernaturals using 
       
     5                   ultrafilters
       
     6 *) 
       
     7        
       
     8 (*------------------------------------------------------------------------
       
     9                        Properties of hypnatrel
       
    10  ------------------------------------------------------------------------*)
       
    11 
       
    12 (** Proving that hyprel is an equivalence relation       **)
       
    13 (** Natural deduction for hypnatrel - similar to hyprel! **)
       
    14 
       
    15 Goalw [hypnatrel_def]
       
    16    "((X,Y): hypnatrel) = ({n. X n = Y n}: FreeUltrafilterNat)";
       
    17 by (Fast_tac 1);
       
    18 qed "hypnatrel_iff";
       
    19 
       
    20 Goalw [hypnatrel_def] 
       
    21      "{n. X n = Y n}: FreeUltrafilterNat ==> (X,Y): hypnatrel";
       
    22 by (Fast_tac 1);
       
    23 qed "hypnatrelI";
       
    24 
       
    25 Goalw [hypnatrel_def]
       
    26   "p: hypnatrel --> (EX X Y. \
       
    27 \                 p = (X,Y) & {n. X n = Y n} : FreeUltrafilterNat)";
       
    28 by (Fast_tac 1);
       
    29 qed "hypnatrelE_lemma";
       
    30 
       
    31 val [major,minor] = goal thy
       
    32   "[| p: hypnatrel;  \
       
    33 \     !!X Y. [| p = (X,Y); {n. X n = Y n}: FreeUltrafilterNat\
       
    34 \            |] ==> Q |] ==> Q";
       
    35 by (cut_facts_tac [major RS (hypnatrelE_lemma RS mp)] 1);
       
    36 by (REPEAT (eresolve_tac [asm_rl,exE,conjE,minor] 1));
       
    37 qed "hypnatrelE";
       
    38 
       
    39 AddSIs [hypnatrelI];
       
    40 AddSEs [hypnatrelE];
       
    41 
       
    42 Goalw [hypnatrel_def] "(x,x): hypnatrel";
       
    43 by (Auto_tac);
       
    44 qed "hypnatrel_refl";
       
    45 
       
    46 Goalw [hypnatrel_def] "(x,y): hypnatrel --> (y,x):hypnatrel";
       
    47 by (auto_tac (claset() addIs [lemma_perm RS subst],simpset()));
       
    48 qed_spec_mp "hypnatrel_sym";
       
    49 
       
    50 Goalw [hypnatrel_def]
       
    51       "(x,y): hypnatrel --> (y,z):hypnatrel --> (x,z):hypnatrel";
       
    52 by (Auto_tac);
       
    53 by (Fuf_tac 1);
       
    54 qed_spec_mp "hypnatrel_trans";
       
    55 
       
    56 Goalw [equiv_def, refl_def, sym_def, trans_def]
       
    57     "equiv {x::nat=>nat. True} hypnatrel";
       
    58 by (auto_tac (claset() addSIs [hypnatrel_refl] addSEs 
       
    59     [hypnatrel_sym,hypnatrel_trans] delrules [hypnatrelI,hypnatrelE],
       
    60     simpset()));
       
    61 qed "equiv_hypnatrel";
       
    62 
       
    63 val equiv_hypnatrel_iff =
       
    64     [TrueI, TrueI] MRS 
       
    65     ([CollectI, CollectI] MRS 
       
    66     (equiv_hypnatrel RS eq_equiv_class_iff));
       
    67 
       
    68 Goalw  [hypnat_def,hypnatrel_def,quotient_def] "hypnatrel^^{x}:hypnat";
       
    69 by (Blast_tac 1);
       
    70 qed "hypnatrel_in_hypnat";
       
    71 
       
    72 Goal "inj_on Abs_hypnat hypnat";
       
    73 by (rtac inj_on_inverseI 1);
       
    74 by (etac Abs_hypnat_inverse 1);
       
    75 qed "inj_on_Abs_hypnat";
       
    76 
       
    77 Addsimps [equiv_hypnatrel_iff,inj_on_Abs_hypnat RS inj_on_iff,
       
    78           hypnatrel_iff, hypnatrel_in_hypnat, Abs_hypnat_inverse];
       
    79 
       
    80 Addsimps [equiv_hypnatrel RS eq_equiv_class_iff];
       
    81 val eq_hypnatrelD = equiv_hypnatrel RSN (2,eq_equiv_class);
       
    82 
       
    83 Goal "inj(Rep_hypnat)";
       
    84 by (rtac inj_inverseI 1);
       
    85 by (rtac Rep_hypnat_inverse 1);
       
    86 qed "inj_Rep_hypnat";
       
    87 
       
    88 Goalw [hypnatrel_def] "x: hypnatrel ^^ {x}";
       
    89 by (Step_tac 1);
       
    90 by (Auto_tac);
       
    91 qed "lemma_hypnatrel_refl";
       
    92 
       
    93 Addsimps [lemma_hypnatrel_refl];
       
    94 
       
    95 Goalw [hypnat_def] "{} ~: hypnat";
       
    96 by (auto_tac (claset() addSEs [quotientE],simpset()));
       
    97 qed "hypnat_empty_not_mem";
       
    98 
       
    99 Addsimps [hypnat_empty_not_mem];
       
   100 
       
   101 Goal "Rep_hypnat x ~= {}";
       
   102 by (cut_inst_tac [("x","x")] Rep_hypnat 1);
       
   103 by (Auto_tac);
       
   104 qed "Rep_hypnat_nonempty";
       
   105 
       
   106 Addsimps [Rep_hypnat_nonempty];
       
   107 
       
   108 (*------------------------------------------------------------------------
       
   109    hypnat_of_nat: the injection from nat to hypnat
       
   110  ------------------------------------------------------------------------*)
       
   111 Goal "inj(hypnat_of_nat)";
       
   112 by (rtac injI 1);
       
   113 by (rewtac hypnat_of_nat_def);
       
   114 by (dtac (inj_on_Abs_hypnat RS inj_onD) 1);
       
   115 by (REPEAT (rtac hypnatrel_in_hypnat 1));
       
   116 by (dtac eq_equiv_class 1);
       
   117 by (rtac equiv_hypnatrel 1);
       
   118 by (Fast_tac 1);
       
   119 by (rtac ccontr 1 THEN rotate_tac 1 1);
       
   120 by (Auto_tac);
       
   121 qed "inj_hypnat_of_nat";
       
   122 
       
   123 val [prem] = goal thy
       
   124     "(!!x. z = Abs_hypnat(hypnatrel^^{x}) ==> P) ==> P";
       
   125 by (res_inst_tac [("x1","z")] 
       
   126     (rewrite_rule [hypnat_def] Rep_hypnat RS quotientE) 1);
       
   127 by (dres_inst_tac [("f","Abs_hypnat")] arg_cong 1);
       
   128 by (res_inst_tac [("x","x")] prem 1);
       
   129 by (asm_full_simp_tac (simpset() addsimps [Rep_hypnat_inverse]) 1);
       
   130 qed "eq_Abs_hypnat";
       
   131 
       
   132 (*-----------------------------------------------------------
       
   133    Addition for hyper naturals: hypnat_add 
       
   134  -----------------------------------------------------------*)
       
   135 Goalw [congruent2_def]
       
   136     "congruent2 hypnatrel (%X Y. hypnatrel^^{%n. X n + Y n})";
       
   137 by Safe_tac;
       
   138 by (ALLGOALS(Fuf_tac));
       
   139 qed "hypnat_add_congruent2";
       
   140 
       
   141 Goalw [hypnat_add_def]
       
   142   "Abs_hypnat(hypnatrel^^{%n. X n}) + Abs_hypnat(hypnatrel^^{%n. Y n}) = \
       
   143 \  Abs_hypnat(hypnatrel^^{%n. X n + Y n})";
       
   144 by (asm_simp_tac
       
   145     (simpset() addsimps [[equiv_hypnatrel, hypnat_add_congruent2] 
       
   146      MRS UN_equiv_class2]) 1);
       
   147 qed "hypnat_add";
       
   148 
       
   149 Goal "(z::hypnat) + w = w + z";
       
   150 by (res_inst_tac [("z","z")] eq_Abs_hypnat 1);
       
   151 by (res_inst_tac [("z","w")] eq_Abs_hypnat 1);
       
   152 by (asm_simp_tac (simpset() addsimps (add_ac @ [hypnat_add])) 1);
       
   153 qed "hypnat_add_commute";
       
   154 
       
   155 Goal "((z1::hypnat) + z2) + z3 = z1 + (z2 + z3)";
       
   156 by (res_inst_tac [("z","z1")] eq_Abs_hypnat 1);
       
   157 by (res_inst_tac [("z","z2")] eq_Abs_hypnat 1);
       
   158 by (res_inst_tac [("z","z3")] eq_Abs_hypnat 1);
       
   159 by (asm_simp_tac (simpset() addsimps [hypnat_add,add_assoc]) 1);
       
   160 qed "hypnat_add_assoc";
       
   161 
       
   162 (*For AC rewriting*)
       
   163 Goal "(x::hypnat)+(y+z)=y+(x+z)";
       
   164 by (rtac (hypnat_add_commute RS trans) 1);
       
   165 by (rtac (hypnat_add_assoc RS trans) 1);
       
   166 by (rtac (hypnat_add_commute RS arg_cong) 1);
       
   167 qed "hypnat_add_left_commute";
       
   168 
       
   169 (* hypnat addition is an AC operator *)
       
   170 val hypnat_add_ac = [hypnat_add_assoc,hypnat_add_commute,
       
   171                       hypnat_add_left_commute];
       
   172 
       
   173 Goalw [hypnat_zero_def] "(0::hypnat) + z = z";
       
   174 by (res_inst_tac [("z","z")] eq_Abs_hypnat 1);
       
   175 by (asm_full_simp_tac (simpset() addsimps 
       
   176     [hypnat_add]) 1);
       
   177 qed "hypnat_add_zero_left";
       
   178 
       
   179 Goal "z + (0::hypnat) = z";
       
   180 by (simp_tac (simpset() addsimps 
       
   181     [hypnat_add_zero_left,hypnat_add_commute]) 1);
       
   182 qed "hypnat_add_zero_right";
       
   183 
       
   184 Addsimps [hypnat_add_zero_left,hypnat_add_zero_right];
       
   185 
       
   186 (*-----------------------------------------------------------
       
   187    Subtraction for hyper naturals: hypnat_minus
       
   188  -----------------------------------------------------------*)
       
   189 Goalw [congruent2_def]
       
   190     "congruent2 hypnatrel (%X Y. hypnatrel^^{%n. X n - Y n})";
       
   191 by Safe_tac;
       
   192 by (ALLGOALS(Fuf_tac));
       
   193 qed "hypnat_minus_congruent2";
       
   194  
       
   195 Goalw [hypnat_minus_def]
       
   196   "Abs_hypnat(hypnatrel^^{%n. X n}) - Abs_hypnat(hypnatrel^^{%n. Y n}) = \
       
   197 \  Abs_hypnat(hypnatrel^^{%n. X n - Y n})";
       
   198 by (asm_simp_tac
       
   199     (simpset() addsimps [[equiv_hypnatrel, hypnat_minus_congruent2] 
       
   200      MRS UN_equiv_class2]) 1);
       
   201 qed "hypnat_minus";
       
   202 
       
   203 Goalw [hypnat_zero_def] "z - z = (0::hypnat)";
       
   204 by (res_inst_tac [("z","z")] eq_Abs_hypnat 1);
       
   205 by (asm_full_simp_tac (simpset() addsimps [hypnat_minus]) 1);
       
   206 qed "hypnat_minus_zero";
       
   207 
       
   208 Goalw [hypnat_zero_def] "(0::hypnat) - n = 0";
       
   209 by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
       
   210 by (asm_full_simp_tac (simpset() addsimps [hypnat_minus]) 1);
       
   211 qed "hypnat_diff_0_eq_0";
       
   212 
       
   213 Addsimps [hypnat_minus_zero,hypnat_diff_0_eq_0];
       
   214 
       
   215 Goalw [hypnat_zero_def] "(m+n = (0::hypnat)) = (m=0 & n=0)";
       
   216 by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
       
   217 by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
       
   218 by (auto_tac (claset() addIs [FreeUltrafilterNat_subset]
       
   219     addDs [FreeUltrafilterNat_Int],
       
   220     simpset() addsimps [hypnat_add] ));
       
   221 qed "hypnat_add_is_0";
       
   222 
       
   223 AddIffs [hypnat_add_is_0];
       
   224 
       
   225 Goal "!!i::hypnat. i-j-k = i - (j+k)";
       
   226 by (res_inst_tac [("z","i")] eq_Abs_hypnat 1);
       
   227 by (res_inst_tac [("z","j")] eq_Abs_hypnat 1);
       
   228 by (res_inst_tac [("z","k")] eq_Abs_hypnat 1);
       
   229 by (asm_full_simp_tac (simpset() addsimps 
       
   230     [hypnat_minus,hypnat_add,diff_diff_left]) 1);
       
   231 qed "hypnat_diff_diff_left";
       
   232 
       
   233 Goal "!!i::hypnat. i-j-k = i-k-j";
       
   234 by (simp_tac (simpset() addsimps 
       
   235     [hypnat_diff_diff_left, hypnat_add_commute]) 1);
       
   236 qed "hypnat_diff_commute";
       
   237 
       
   238 Goal "!!n::hypnat. (n+m) - n = m";
       
   239 by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
       
   240 by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
       
   241 by (asm_full_simp_tac (simpset() addsimps 
       
   242     [hypnat_minus,hypnat_add]) 1);
       
   243 qed "hypnat_diff_add_inverse";
       
   244 Addsimps [hypnat_diff_add_inverse];
       
   245 
       
   246 Goal "!!n::hypnat.(m+n) - n = m";
       
   247 by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
       
   248 by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
       
   249 by (asm_full_simp_tac (simpset() addsimps 
       
   250     [hypnat_minus,hypnat_add]) 1);
       
   251 qed "hypnat_diff_add_inverse2";
       
   252 Addsimps [hypnat_diff_add_inverse2];
       
   253 
       
   254 Goal "!!k::hypnat. (k+m) - (k+n) = m - n";
       
   255 by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
       
   256 by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
       
   257 by (res_inst_tac [("z","k")] eq_Abs_hypnat 1);
       
   258 by (asm_full_simp_tac (simpset() addsimps 
       
   259     [hypnat_minus,hypnat_add]) 1);
       
   260 qed "hypnat_diff_cancel";
       
   261 Addsimps [hypnat_diff_cancel];
       
   262 
       
   263 Goal "!!m::hypnat. (m+k) - (n+k) = m - n";
       
   264 val hypnat_add_commute_k = read_instantiate [("w","k")] hypnat_add_commute;
       
   265 by (asm_simp_tac (simpset() addsimps ([hypnat_add_commute_k])) 1);
       
   266 qed "hypnat_diff_cancel2";
       
   267 Addsimps [hypnat_diff_cancel2];
       
   268 
       
   269 Goalw [hypnat_zero_def] "!!n::hypnat. n - (n+m) = (0::hypnat)";
       
   270 by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
       
   271 by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
       
   272 by (asm_full_simp_tac (simpset() addsimps 
       
   273     [hypnat_minus,hypnat_add]) 1);
       
   274 qed "hypnat_diff_add_0";
       
   275 Addsimps [hypnat_diff_add_0];
       
   276 
       
   277 (*-----------------------------------------------------------
       
   278    Multiplication for hyper naturals: hypnat_mult
       
   279  -----------------------------------------------------------*)
       
   280 Goalw [congruent2_def]
       
   281     "congruent2 hypnatrel (%X Y. hypnatrel^^{%n. X n * Y n})";
       
   282 by Safe_tac;
       
   283 by (ALLGOALS(Fuf_tac));
       
   284 qed "hypnat_mult_congruent2";
       
   285 
       
   286 Goalw [hypnat_mult_def]
       
   287   "Abs_hypnat(hypnatrel^^{%n. X n}) * Abs_hypnat(hypnatrel^^{%n. Y n}) = \
       
   288 \  Abs_hypnat(hypnatrel^^{%n. X n * Y n})";
       
   289 by (asm_simp_tac
       
   290     (simpset() addsimps [[equiv_hypnatrel,hypnat_mult_congruent2] MRS
       
   291      UN_equiv_class2]) 1);
       
   292 qed "hypnat_mult";
       
   293 
       
   294 Goal "(z::hypnat) * w = w * z";
       
   295 by (res_inst_tac [("z","z")] eq_Abs_hypnat 1);
       
   296 by (res_inst_tac [("z","w")] eq_Abs_hypnat 1);
       
   297 by (asm_simp_tac (simpset() addsimps ([hypnat_mult] @ mult_ac)) 1);
       
   298 qed "hypnat_mult_commute";
       
   299 
       
   300 Goal "((z1::hypnat) * z2) * z3 = z1 * (z2 * z3)";
       
   301 by (res_inst_tac [("z","z1")] eq_Abs_hypnat 1);
       
   302 by (res_inst_tac [("z","z2")] eq_Abs_hypnat 1);
       
   303 by (res_inst_tac [("z","z3")] eq_Abs_hypnat 1);
       
   304 by (asm_simp_tac (simpset() addsimps [hypnat_mult,mult_assoc]) 1);
       
   305 qed "hypnat_mult_assoc";
       
   306 
       
   307 qed_goal "hypnat_mult_left_commute" thy
       
   308     "(z1::hypnat) * (z2 * z3) = z2 * (z1 * z3)"
       
   309  (fn _ => [rtac (hypnat_mult_commute RS trans) 1, rtac (hypnat_mult_assoc RS trans) 1,
       
   310            rtac (hypnat_mult_commute RS arg_cong) 1]);
       
   311 
       
   312 (* hypnat multiplication is an AC operator *)
       
   313 val hypnat_mult_ac = [hypnat_mult_assoc, hypnat_mult_commute, 
       
   314                        hypnat_mult_left_commute];
       
   315 
       
   316 Goalw [hypnat_one_def] "1hn * z = z";
       
   317 by (res_inst_tac [("z","z")] eq_Abs_hypnat 1);
       
   318 by (asm_full_simp_tac (simpset() addsimps [hypnat_mult]) 1);
       
   319 qed "hypnat_mult_1";
       
   320 Addsimps [hypnat_mult_1];
       
   321 
       
   322 Goal "z * 1hn = z";
       
   323 by (simp_tac (simpset() addsimps [hypnat_mult_commute]) 1);
       
   324 qed "hypnat_mult_1_right";
       
   325 Addsimps [hypnat_mult_1_right];
       
   326 
       
   327 Goalw [hypnat_zero_def] "(0::hypnat) * z = 0";
       
   328 by (res_inst_tac [("z","z")] eq_Abs_hypnat 1);
       
   329 by (asm_full_simp_tac (simpset() addsimps [hypnat_mult]) 1);
       
   330 qed "hypnat_mult_0";
       
   331 Addsimps [hypnat_mult_0];
       
   332 
       
   333 Goal "z * (0::hypnat) = 0";
       
   334 by (simp_tac (simpset() addsimps [hypnat_mult_commute]) 1);
       
   335 qed "hypnat_mult_0_right";
       
   336 Addsimps [hypnat_mult_0_right];
       
   337 
       
   338 Goal "!!m::hypnat. (m - n) * k = (m * k) - (n * k)";
       
   339 by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
       
   340 by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
       
   341 by (res_inst_tac [("z","k")] eq_Abs_hypnat 1);
       
   342 by (asm_simp_tac (simpset() addsimps [hypnat_mult,
       
   343     hypnat_minus,diff_mult_distrib]) 1);
       
   344 qed "hypnat_diff_mult_distrib" ;
       
   345 
       
   346 Goal "!!m::hypnat. k * (m - n) = (k * m) - (k * n)";
       
   347 val hypnat_mult_commute_k = read_instantiate [("z","k")] hypnat_mult_commute;
       
   348 by (simp_tac (simpset() addsimps [hypnat_diff_mult_distrib, 
       
   349     hypnat_mult_commute_k]) 1);
       
   350 qed "hypnat_diff_mult_distrib2" ;
       
   351 
       
   352 (*--------------------------
       
   353     A Few more theorems
       
   354  -------------------------*)
       
   355 qed_goal "hypnat_add_assoc_cong" thy
       
   356     "!!z. (z::hypnat) + v = z' + v' ==> z + (v + w) = z' + (v' + w)"
       
   357  (fn _ => [(asm_simp_tac (simpset() addsimps [hypnat_add_assoc RS sym]) 1)]);
       
   358 
       
   359 qed_goal "hypnat_add_assoc_swap" thy 
       
   360           "(z::hypnat) + (v + w) = v + (z + w)"
       
   361  (fn _ => [(REPEAT (ares_tac [hypnat_add_commute RS hypnat_add_assoc_cong] 1))]);
       
   362 
       
   363 Goal "((z1::hypnat) + z2) * w = (z1 * w) + (z2 * w)";
       
   364 by (res_inst_tac [("z","z1")] eq_Abs_hypnat 1);
       
   365 by (res_inst_tac [("z","z2")] eq_Abs_hypnat 1);
       
   366 by (res_inst_tac [("z","w")] eq_Abs_hypnat 1);
       
   367 by (asm_simp_tac (simpset() addsimps [hypnat_mult,hypnat_add,
       
   368      add_mult_distrib]) 1);
       
   369 qed "hypnat_add_mult_distrib";
       
   370 Addsimps [hypnat_add_mult_distrib];
       
   371 
       
   372 val hypnat_mult_commute'= read_instantiate [("z","w")] hypnat_mult_commute;
       
   373 
       
   374 Goal "(w::hypnat) * (z1 + z2) = (w * z1) + (w * z2)";
       
   375 by (simp_tac (simpset() addsimps [hypnat_mult_commute']) 1);
       
   376 qed "hypnat_add_mult_distrib2";
       
   377 Addsimps [hypnat_add_mult_distrib2];
       
   378 
       
   379 (*** (hypnat) one and zero are distinct ***)
       
   380 Goalw [hypnat_zero_def,hypnat_one_def] "(0::hypnat) ~= 1hn";
       
   381 by (Auto_tac);
       
   382 qed "hypnat_zero_not_eq_one";
       
   383 Addsimps [hypnat_zero_not_eq_one];
       
   384 Addsimps [hypnat_zero_not_eq_one RS not_sym];
       
   385 
       
   386 Goalw [hypnat_zero_def] "(m*n = (0::hypnat)) = (m=0 | n=0)";
       
   387 by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
       
   388 by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
       
   389 by (auto_tac (claset() addSDs [FreeUltrafilterNat_Compl_mem],
       
   390     simpset() addsimps [hypnat_mult]));
       
   391 by (ALLGOALS(Fuf_tac));
       
   392 qed "hypnat_mult_is_0";
       
   393 Addsimps [hypnat_mult_is_0];
       
   394 
       
   395 (*------------------------------------------------------------------
       
   396                    Theorems for ordering 
       
   397  ------------------------------------------------------------------*)
       
   398 
       
   399 (* prove introduction and elimination rules for hypnat_less *)
       
   400 
       
   401 Goalw [hypnat_less_def]
       
   402  "P < (Q::hypnat) = (EX X Y. X : Rep_hypnat(P) & \
       
   403 \                             Y : Rep_hypnat(Q) & \
       
   404 \                             {n. X n < Y n} : FreeUltrafilterNat)";
       
   405 by (Fast_tac 1);
       
   406 qed "hypnat_less_iff";
       
   407 
       
   408 Goalw [hypnat_less_def]
       
   409  "!!P. [| {n. X n < Y n} : FreeUltrafilterNat; \
       
   410 \         X : Rep_hypnat(P); \
       
   411 \         Y : Rep_hypnat(Q) |] ==> P < (Q::hypnat)";
       
   412 by (Fast_tac 1);
       
   413 qed "hypnat_lessI";
       
   414 
       
   415 Goalw [hypnat_less_def]
       
   416      "!! R1. [| R1 < (R2::hypnat); \
       
   417 \         !!X Y. {n. X n < Y n} : FreeUltrafilterNat ==> P; \
       
   418 \         !!X. X : Rep_hypnat(R1) ==> P; \ 
       
   419 \         !!Y. Y : Rep_hypnat(R2) ==> P |] \
       
   420 \     ==> P";
       
   421 by (Auto_tac);
       
   422 qed "hypnat_lessE";
       
   423 
       
   424 Goalw [hypnat_less_def]
       
   425  "!!R1. R1 < (R2::hypnat) ==> (EX X Y. {n. X n < Y n} : FreeUltrafilterNat & \
       
   426 \                                  X : Rep_hypnat(R1) & \
       
   427 \                                  Y : Rep_hypnat(R2))";
       
   428 by (Fast_tac 1);
       
   429 qed "hypnat_lessD";
       
   430 
       
   431 Goal "~ (R::hypnat) < R";
       
   432 by (res_inst_tac [("z","R")] eq_Abs_hypnat 1);
       
   433 by (auto_tac (claset(),simpset() addsimps [hypnat_less_def]));
       
   434 by (Fuf_empty_tac 1);
       
   435 qed "hypnat_less_not_refl";
       
   436 Addsimps [hypnat_less_not_refl];
       
   437 
       
   438 bind_thm("hypnat_less_irrefl",hypnat_less_not_refl RS notE);
       
   439 
       
   440 qed_goal "hypnat_not_refl2" thy 
       
   441     "!!(x::hypnat). x < y ==> x ~= y" (fn _ => [Auto_tac]);
       
   442 
       
   443 Goalw [hypnat_less_def,hypnat_zero_def] "~ n<(0::hypnat)";
       
   444 by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
       
   445 by (Auto_tac);
       
   446 by (Fuf_empty_tac 1);
       
   447 qed "hypnat_not_less0";
       
   448 AddIffs [hypnat_not_less0];
       
   449 
       
   450 (* n<(0::hypnat) ==> R *)
       
   451 bind_thm ("hypnat_less_zeroE", hypnat_not_less0 RS notE);
       
   452 
       
   453 Goalw [hypnat_less_def,hypnat_zero_def,hypnat_one_def]
       
   454       "(n<1hn) = (n=0)";
       
   455 by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
       
   456 by (auto_tac (claset() addSIs [exI] addEs 
       
   457     [FreeUltrafilterNat_subset],simpset()));
       
   458 by (Fuf_tac 1);
       
   459 qed "hypnat_less_one";
       
   460 AddIffs [hypnat_less_one];
       
   461 
       
   462 Goal "!!(R1::hypnat). [| R1 < R2; R2 < R3 |] ==> R1 < R3";
       
   463 by (res_inst_tac [("z","R1")] eq_Abs_hypnat 1);
       
   464 by (res_inst_tac [("z","R2")] eq_Abs_hypnat 1);
       
   465 by (res_inst_tac [("z","R3")] eq_Abs_hypnat 1);
       
   466 by (auto_tac (claset(),simpset() addsimps [hypnat_less_def]));
       
   467 by (res_inst_tac [("x","X")] exI 1);
       
   468 by (Auto_tac);
       
   469 by (res_inst_tac [("x","Ya")] exI 1);
       
   470 by (Auto_tac);
       
   471 by (Fuf_tac 1);
       
   472 qed "hypnat_less_trans";
       
   473 
       
   474 Goal "!! (R1::hypnat). [| R1 < R2; R2 < R1 |] ==> P";
       
   475 by (dtac hypnat_less_trans 1 THEN assume_tac 1);
       
   476 by (Asm_full_simp_tac 1);
       
   477 qed "hypnat_less_asym";
       
   478 
       
   479 (*----- hypnat less iff less a.e -----*)
       
   480 (* See comments in HYPER for corresponding thm *)
       
   481 
       
   482 Goalw [hypnat_less_def]
       
   483       "(Abs_hypnat(hypnatrel^^{%n. X n}) < \
       
   484 \           Abs_hypnat(hypnatrel^^{%n. Y n})) = \
       
   485 \      ({n. X n < Y n} : FreeUltrafilterNat)";
       
   486 by (auto_tac (claset() addSIs [lemma_hypnatrel_refl],simpset()));
       
   487 by (Fuf_tac 1);
       
   488 qed "hypnat_less";
       
   489 
       
   490 Goal "~ m<n --> n+(m-n) = (m::hypnat)";
       
   491 by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
       
   492 by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
       
   493 by (auto_tac (claset(),simpset() addsimps 
       
   494     [hypnat_minus,hypnat_add,hypnat_less]));
       
   495 by (dtac FreeUltrafilterNat_Compl_mem 1);
       
   496 by (Fuf_tac 1);
       
   497 qed_spec_mp "hypnat_add_diff_inverse";
       
   498 
       
   499 Goal "n<=m ==> n+(m-n) = (m::hypnat)";
       
   500 by (asm_full_simp_tac (simpset() addsimps 
       
   501     [hypnat_add_diff_inverse, hypnat_le_def]) 1);
       
   502 qed "hypnat_le_add_diff_inverse";
       
   503 
       
   504 Goal "n<=m ==> (m-n)+n = (m::hypnat)";
       
   505 by (asm_simp_tac (simpset() addsimps [hypnat_le_add_diff_inverse, 
       
   506     hypnat_add_commute]) 1);
       
   507 qed "hypnat_le_add_diff_inverse2";
       
   508 
       
   509 (*---------------------------------------------------------------------------------
       
   510                     Hyper naturals as a linearly ordered field
       
   511  ---------------------------------------------------------------------------------*)
       
   512 Goalw [hypnat_zero_def] 
       
   513      "[| (0::hypnat) < x; 0 < y |] ==> 0 < x + y";
       
   514 by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
       
   515 by (res_inst_tac [("z","y")] eq_Abs_hypnat 1);
       
   516 by (auto_tac (claset(),simpset() addsimps
       
   517    [hypnat_less_def,hypnat_add]));
       
   518 by (REPEAT(Step_tac 1));
       
   519 by (Fuf_tac 1);
       
   520 qed "hypnat_add_order";
       
   521 
       
   522 Goalw [hypnat_zero_def] 
       
   523       "!!(x::hypnat). [| (0::hypnat) < x; 0 < y |] ==> 0 < x * y";
       
   524 by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
       
   525 by (res_inst_tac [("z","y")] eq_Abs_hypnat 1);
       
   526 by (auto_tac (claset(),simpset() addsimps 
       
   527     [hypnat_less_def,hypnat_mult]));
       
   528 by (REPEAT(Step_tac 1));
       
   529 by (Fuf_tac 1);
       
   530 qed "hypnat_mult_order";
       
   531 
       
   532 (*---------------------------------------------------------------------------------
       
   533                    Trichotomy of the hyper naturals
       
   534   --------------------------------------------------------------------------------*)
       
   535 Goalw [hypnatrel_def] "EX x. x: hypnatrel ^^ {%n. 0}";
       
   536 by (res_inst_tac [("x","%n. 0")] exI 1);
       
   537 by (Step_tac 1);
       
   538 by (Auto_tac);
       
   539 qed "lemma_hypnatrel_0_mem";
       
   540 
       
   541 (* linearity is actually proved further down! *)
       
   542 Goalw [hypnat_zero_def,
       
   543        hypnat_less_def]"(0::hypnat) <  x | x = 0 | x < 0";
       
   544 by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
       
   545 by (Auto_tac );
       
   546 by (REPEAT(Step_tac 1));
       
   547 by (REPEAT(dtac FreeUltrafilterNat_Compl_mem 1));
       
   548 by (Fuf_tac 1);
       
   549 qed "hypnat_trichotomy";
       
   550 
       
   551 Goal "!!x. [| (0::hypnat) < x ==> P; \
       
   552 \                 x = 0 ==> P; \
       
   553 \                 x < 0 ==> P |] ==> P";
       
   554 by (cut_inst_tac [("x","x")] hypnat_trichotomy 1);
       
   555 by (Auto_tac);
       
   556 qed "hypnat_trichotomyE";
       
   557 
       
   558 (*------------------------------------------------------------------------------
       
   559             More properties of <
       
   560  ------------------------------------------------------------------------------*)
       
   561 Goal "!!(A::hypnat). A < B ==> A + C < B + C";
       
   562 by (res_inst_tac [("z","A")] eq_Abs_hypnat 1);
       
   563 by (res_inst_tac [("z","B")] eq_Abs_hypnat 1);
       
   564 by (res_inst_tac [("z","C")] eq_Abs_hypnat 1);
       
   565 by (auto_tac (claset(),simpset() addsimps 
       
   566     [hypnat_less_def,hypnat_add]));
       
   567 by (REPEAT(Step_tac 1));
       
   568 by (Fuf_tac 1);
       
   569 qed "hypnat_add_less_mono1";
       
   570 
       
   571 Goal "!!(A::hypnat). A < B ==> C + A < C + B";
       
   572 by (auto_tac (claset() addIs [hypnat_add_less_mono1],
       
   573     simpset() addsimps [hypnat_add_commute]));
       
   574 qed "hypnat_add_less_mono2";
       
   575 
       
   576 Goal "!!k l::hypnat. [|i<j;  k<l |] ==> i + k < j + l";
       
   577 by (etac (hypnat_add_less_mono1 RS hypnat_less_trans) 1);
       
   578 by (simp_tac (simpset() addsimps [hypnat_add_commute]) 1);
       
   579 (*j moves to the end because it is free while k, l are bound*)
       
   580 by (etac hypnat_add_less_mono1 1);
       
   581 qed "hypnat_add_less_mono";
       
   582 
       
   583 (*---------------------------------------
       
   584         hypnat_minus_less
       
   585  ---------------------------------------*)
       
   586 Goalw [hypnat_less_def,hypnat_zero_def] 
       
   587       "((x::hypnat) < y) = ((0::hypnat) < y - x)";
       
   588 by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
       
   589 by (res_inst_tac [("z","y")] eq_Abs_hypnat 1);
       
   590 by (auto_tac (claset(),simpset() addsimps 
       
   591     [hypnat_minus]));
       
   592 by (REPEAT(Step_tac 1));
       
   593 by (REPEAT(Step_tac 2));
       
   594 by (ALLGOALS(fuf_tac (claset() addDs [sym],simpset())));
       
   595 
       
   596 (*** linearity ***)
       
   597 Goalw [hypnat_less_def] "(x::hypnat) < y | x = y | y < x";
       
   598 by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
       
   599 by (res_inst_tac [("z","y")] eq_Abs_hypnat 1);
       
   600 by (Auto_tac );
       
   601 by (REPEAT(Step_tac 1));
       
   602 by (REPEAT(dtac FreeUltrafilterNat_Compl_mem 1));
       
   603 by (Fuf_tac 1);
       
   604 qed "hypnat_linear";
       
   605 
       
   606 Goal
       
   607     "!!(x::hypnat). [| x < y ==> P;  x = y ==> P; \
       
   608 \          y < x ==> P |] ==> P";
       
   609 by (cut_inst_tac [("x","x"),("y","y")] hypnat_linear 1);
       
   610 by (Auto_tac);
       
   611 qed "hypnat_linear_less2";
       
   612 
       
   613 (*------------------------------------------------------------------------------
       
   614                             Properties of <=
       
   615  ------------------------------------------------------------------------------*)
       
   616 (*------ hypnat le iff nat le a.e ------*)
       
   617 Goalw [hypnat_le_def,le_def]
       
   618       "(Abs_hypnat(hypnatrel^^{%n. X n}) <= \
       
   619 \           Abs_hypnat(hypnatrel^^{%n. Y n})) = \
       
   620 \      ({n. X n <= Y n} : FreeUltrafilterNat)";
       
   621 by (auto_tac (claset() addSDs [FreeUltrafilterNat_Compl_mem],
       
   622     simpset() addsimps [hypnat_less]));
       
   623 by (Fuf_tac 1 THEN Fuf_empty_tac 1);
       
   624 qed "hypnat_le";
       
   625 
       
   626 (*---------------------------------------------------------*)
       
   627 (*---------------------------------------------------------*)
       
   628 Goalw [hypnat_le_def] "!!w. ~(w < z) ==> z <= (w::hypnat)";
       
   629 by (assume_tac 1);
       
   630 qed "hypnat_leI";
       
   631 
       
   632 Goalw [hypnat_le_def] "!!w. z<=w ==> ~(w<(z::hypnat))";
       
   633 by (assume_tac 1);
       
   634 qed "hypnat_leD";
       
   635 
       
   636 val hypnat_leE = make_elim hypnat_leD;
       
   637 
       
   638 Goal "!!w. (~(w < z)) = (z <= (w::hypnat))";
       
   639 by (fast_tac (claset() addSIs [hypnat_leI,hypnat_leD]) 1);
       
   640 qed "hypnat_less_le_iff";
       
   641 
       
   642 Goalw [hypnat_le_def] "!!z. ~ z <= w ==> w<(z::hypnat)";
       
   643 by (Fast_tac 1);
       
   644 qed "not_hypnat_leE";
       
   645 
       
   646 Goalw [hypnat_le_def] "!!z. z < w ==> z <= (w::hypnat)";
       
   647 by (fast_tac (claset() addEs [hypnat_less_asym]) 1);
       
   648 qed "hypnat_less_imp_le";
       
   649 
       
   650 Goalw [hypnat_le_def] "!!(x::hypnat). x <= y ==> x < y | x = y";
       
   651 by (cut_facts_tac [hypnat_linear] 1);
       
   652 by (fast_tac (claset() addEs [hypnat_less_irrefl,hypnat_less_asym]) 1);
       
   653 qed "hypnat_le_imp_less_or_eq";
       
   654 
       
   655 Goalw [hypnat_le_def] "!!z. z<w | z=w ==> z <=(w::hypnat)";
       
   656 by (cut_facts_tac [hypnat_linear] 1);
       
   657 by (fast_tac (claset() addEs [hypnat_less_irrefl,hypnat_less_asym]) 1);
       
   658 qed "hypnat_less_or_eq_imp_le";
       
   659 
       
   660 Goal "(x <= (y::hypnat)) = (x < y | x=y)";
       
   661 by (REPEAT(ares_tac [iffI, hypnat_less_or_eq_imp_le, hypnat_le_imp_less_or_eq] 1));
       
   662 qed "hypnat_le_eq_less_or_eq";
       
   663 
       
   664 Goal "w <= (w::hypnat)";
       
   665 by (simp_tac (simpset() addsimps [hypnat_le_eq_less_or_eq]) 1);
       
   666 qed "hypnat_le_refl";
       
   667 Addsimps [hypnat_le_refl];
       
   668 
       
   669 val prems = goal thy "!!i. [| i <= j; j < k |] ==> i < (k::hypnat)";
       
   670 by (dtac hypnat_le_imp_less_or_eq 1);
       
   671 by (fast_tac (claset() addIs [hypnat_less_trans]) 1);
       
   672 qed "hypnat_le_less_trans";
       
   673 
       
   674 Goal "!! (i::hypnat). [| i < j; j <= k |] ==> i < k";
       
   675 by (dtac hypnat_le_imp_less_or_eq 1);
       
   676 by (fast_tac (claset() addIs [hypnat_less_trans]) 1);
       
   677 qed "hypnat_less_le_trans";
       
   678 
       
   679 Goal "!!i. [| i <= j; j <= k |] ==> i <= (k::hypnat)";
       
   680 by (EVERY1 [dtac hypnat_le_imp_less_or_eq, dtac hypnat_le_imp_less_or_eq,
       
   681             rtac hypnat_less_or_eq_imp_le, fast_tac (claset() addIs [hypnat_less_trans])]);
       
   682 qed "hypnat_le_trans";
       
   683 
       
   684 Goal "!!z. [| z <= w; w <= z |] ==> z = (w::hypnat)";
       
   685 by (EVERY1 [dtac hypnat_le_imp_less_or_eq, dtac hypnat_le_imp_less_or_eq,
       
   686             fast_tac (claset() addEs [hypnat_less_irrefl,hypnat_less_asym])]);
       
   687 qed "hypnat_le_anti_sym";
       
   688 
       
   689 Goal "!!x. [| ~ y < x; y ~= x |] ==> x < (y::hypnat)";
       
   690 by (rtac not_hypnat_leE 1);
       
   691 by (fast_tac (claset() addDs [hypnat_le_imp_less_or_eq]) 1);
       
   692 qed "not_less_not_eq_hypnat_less";
       
   693 
       
   694 Goal "!!x. [| (0::hypnat) <= x; 0 <= y |] ==> 0 <= x * y";
       
   695 by (REPEAT(dtac hypnat_le_imp_less_or_eq 1));
       
   696 by (auto_tac (claset() addIs [hypnat_mult_order,
       
   697     hypnat_less_imp_le],simpset() addsimps [hypnat_le_refl]));
       
   698 qed "hypnat_le_mult_order";
       
   699 
       
   700 Goalw [hypnat_one_def,hypnat_zero_def,hypnat_less_def] 
       
   701       "(0::hypnat) < 1hn";
       
   702 by (res_inst_tac [("x","%n. 0")] exI 1);
       
   703 by (res_inst_tac [("x","%n. 1")] exI 1);
       
   704 by (Auto_tac);
       
   705 qed "hypnat_zero_less_one";
       
   706 
       
   707 Goal "!!x. [| (0::hypnat) <= x; 0 <= y |] ==> 0 <= x + y";
       
   708 by (REPEAT(dtac hypnat_le_imp_less_or_eq 1));
       
   709 by (auto_tac (claset() addIs [hypnat_add_order,
       
   710     hypnat_less_imp_le],simpset() addsimps [hypnat_le_refl]));
       
   711 qed "hypnat_le_add_order";
       
   712 
       
   713 Goal "!!(q1::hypnat). q1 <= q2  ==> x + q1 <= x + q2";
       
   714 by (dtac hypnat_le_imp_less_or_eq 1);
       
   715 by (Step_tac 1);
       
   716 by (auto_tac (claset() addSIs [hypnat_le_refl,
       
   717     hypnat_less_imp_le,hypnat_add_less_mono1],
       
   718     simpset() addsimps [hypnat_add_commute]));
       
   719 qed "hypnat_add_le_mono2";
       
   720 
       
   721 Goal "!!(q1::hypnat). q1 <= q2  ==> q1 + x <= q2 + x";
       
   722 by (auto_tac (claset() addDs [hypnat_add_le_mono2],
       
   723     simpset() addsimps [hypnat_add_commute]));
       
   724 qed "hypnat_add_le_mono1";
       
   725 
       
   726 Goal "!!k l::hypnat. [|i<=j;  k<=l |] ==> i + k <= j + l";
       
   727 by (etac (hypnat_add_le_mono1 RS hypnat_le_trans) 1);
       
   728 by (simp_tac (simpset() addsimps [hypnat_add_commute]) 1);
       
   729 (*j moves to the end because it is free while k, l are bound*)
       
   730 by (etac hypnat_add_le_mono1 1);
       
   731 qed "hypnat_add_le_mono";
       
   732 
       
   733 Goalw [hypnat_zero_def]
       
   734      "!!x::hypnat. [| (0::hypnat) < z; x < y |] ==> x * z < y * z";
       
   735 by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
       
   736 by (res_inst_tac [("z","y")] eq_Abs_hypnat 1);
       
   737 by (res_inst_tac [("z","z")] eq_Abs_hypnat 1);
       
   738 by (auto_tac (claset(),simpset() addsimps 
       
   739     [hypnat_less,hypnat_mult]));
       
   740 by (Fuf_tac 1);
       
   741 qed "hypnat_mult_less_mono1";
       
   742 
       
   743 Goal "!!x::hypnat. [| 0 < z; x < y |] ==> z * x < z * y";
       
   744 by (auto_tac (claset() addIs [hypnat_mult_less_mono1],
       
   745               simpset() addsimps [hypnat_mult_commute]));
       
   746 qed "hypnat_mult_less_mono2";
       
   747 
       
   748 Addsimps [hypnat_mult_less_mono2,hypnat_mult_less_mono1];
       
   749 
       
   750 Goal "(x::hypnat) <= n + x";
       
   751 by (res_inst_tac [("x","n")] hypnat_trichotomyE 1);
       
   752 by (auto_tac (claset() addDs [(hypnat_less_imp_le RS
       
   753     hypnat_add_le_mono1)],simpset() addsimps [hypnat_le_refl]));
       
   754 qed "hypnat_add_self_le";
       
   755 Addsimps [hypnat_add_self_le];
       
   756 
       
   757 Goal "(x::hypnat) < x + 1hn";
       
   758 by (cut_facts_tac [hypnat_zero_less_one 
       
   759          RS hypnat_add_less_mono2] 1);
       
   760 by (Auto_tac);
       
   761 qed "hypnat_add_one_self_less";
       
   762 Addsimps [hypnat_add_one_self_less];
       
   763 
       
   764 Goalw [hypnat_le_def] "~ x + 1hn <= x";
       
   765 by (Simp_tac 1);
       
   766 qed "not_hypnat_add_one_le_self";
       
   767 Addsimps [not_hypnat_add_one_le_self];
       
   768 
       
   769 Goal "((0::hypnat) < n) = (1hn <= n)";
       
   770 by (res_inst_tac [("x","n")] hypnat_trichotomyE 1);
       
   771 by (auto_tac (claset(),simpset() addsimps [hypnat_le_def]));
       
   772 qed "hypnat_gt_zero_iff";
       
   773 
       
   774 Addsimps  [hypnat_le_add_diff_inverse, hypnat_le_add_diff_inverse2,
       
   775            hypnat_less_imp_le RS hypnat_le_add_diff_inverse2];
       
   776 
       
   777 Goal "(0 < n) = (EX m. n = m + 1hn)";
       
   778 by (Step_tac 1);
       
   779 by (res_inst_tac [("x","m")] hypnat_trichotomyE 2);
       
   780 by (rtac hypnat_less_trans 2);
       
   781 by (res_inst_tac [("x","n - 1hn")] exI 1);
       
   782 by (auto_tac (claset(),simpset() addsimps 
       
   783     [hypnat_gt_zero_iff,hypnat_add_commute]));
       
   784 qed "hypnat_gt_zero_iff2";
       
   785 
       
   786 Goalw [hypnat_zero_def] "(0::hypnat) <= n";
       
   787 by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
       
   788 by (asm_simp_tac (simpset() addsimps [hypnat_le]) 1);
       
   789 qed "hypnat_le_zero";
       
   790 Addsimps [hypnat_le_zero];
       
   791 
       
   792 (*------------------------------------------------------------------
       
   793      hypnat_of_nat: properties embedding of naturals in hypernaturals
       
   794  -----------------------------------------------------------------*)
       
   795     (** hypnat_of_nat preserves field and order properties **)
       
   796 
       
   797 Goalw [hypnat_of_nat_def] 
       
   798       "hypnat_of_nat ((z1::nat) + z2) = \
       
   799 \      hypnat_of_nat z1 + hypnat_of_nat z2";
       
   800 by (asm_simp_tac (simpset() addsimps [hypnat_add]) 1);
       
   801 qed "hypnat_of_nat_add";
       
   802 
       
   803 Goalw [hypnat_of_nat_def] 
       
   804       "hypnat_of_nat ((z1::nat) - z2) = \
       
   805 \      hypnat_of_nat z1 - hypnat_of_nat z2";
       
   806 by (asm_simp_tac (simpset() addsimps [hypnat_minus]) 1);
       
   807 qed "hypnat_of_nat_minus";
       
   808 
       
   809 Goalw [hypnat_of_nat_def] 
       
   810       "hypnat_of_nat (z1 * z2) = hypnat_of_nat z1 * hypnat_of_nat z2";
       
   811 by (full_simp_tac (simpset() addsimps [hypnat_mult]) 1);
       
   812 qed "hypnat_of_nat_mult";
       
   813 
       
   814 Goalw [hypnat_less_def,hypnat_of_nat_def] 
       
   815       "(z1 < z2) = (hypnat_of_nat z1 < hypnat_of_nat z2)";
       
   816 by (auto_tac (claset() addSIs [exI] addIs 
       
   817    [FreeUltrafilterNat_all],simpset()));
       
   818 by (rtac FreeUltrafilterNat_P 1 THEN Fuf_tac 1);
       
   819 qed "hypnat_of_nat_less_iff";
       
   820 Addsimps [hypnat_of_nat_less_iff RS sym];
       
   821 
       
   822 Goalw [hypnat_le_def,le_def] 
       
   823       "(z1 <= z2) = (hypnat_of_nat z1 <= hypnat_of_nat z2)";
       
   824 by (Auto_tac);
       
   825 qed "hypnat_of_nat_le_iff";
       
   826 
       
   827 Goalw [hypnat_of_nat_def,hypnat_one_def] "hypnat_of_nat  1 = 1hn";
       
   828 by (Simp_tac 1);
       
   829 qed "hypnat_of_nat_one";
       
   830 
       
   831 Goalw [hypnat_of_nat_def,hypnat_zero_def] "hypnat_of_nat  0 = 0";
       
   832 by (Simp_tac 1);
       
   833 qed "hypnat_of_nat_zero";
       
   834 
       
   835 Goal "(hypnat_of_nat  n = 0) = (n = 0)";
       
   836 by (auto_tac (claset() addIs [FreeUltrafilterNat_P],
       
   837     simpset() addsimps [hypnat_of_nat_def,
       
   838     hypnat_zero_def]));
       
   839 qed "hypnat_of_nat_zero_iff";
       
   840 
       
   841 Goal "(hypnat_of_nat  n ~= 0) = (n ~= 0)";
       
   842 by (full_simp_tac (simpset() addsimps [hypnat_of_nat_zero_iff]) 1);
       
   843 qed "hypnat_of_nat_not_zero_iff";
       
   844 
       
   845 goalw HyperNat.thy [hypnat_of_nat_def,hypnat_one_def]
       
   846       "hypnat_of_nat (Suc n) = hypnat_of_nat n + 1hn";
       
   847 by (auto_tac (claset(),simpset() addsimps [hypnat_add]));
       
   848 qed "hypnat_of_nat_Suc";
       
   849 
       
   850 (*---------------------------------------------------------------------------------
       
   851               Existence of infinite hypernatural number
       
   852  ---------------------------------------------------------------------------------*)
       
   853 
       
   854 Goal "hypnatrel^^{%n::nat. n} : hypnat";
       
   855 by (Auto_tac);
       
   856 qed "hypnat_omega";
       
   857 
       
   858 Goalw [hypnat_omega_def] "Rep_hypnat(whn) : hypnat";
       
   859 by (rtac Rep_hypnat 1);
       
   860 qed "Rep_hypnat_omega";
       
   861 
       
   862 (* See Hyper.thy for similar argument*)
       
   863 (* existence of infinite number not corresponding to any natural number *)
       
   864 (* use assumption that member FreeUltrafilterNat is not finite       *)
       
   865 (* a few lemmas first *)
       
   866 
       
   867 Goalw [hypnat_omega_def,hypnat_of_nat_def]
       
   868       "~ (EX x. hypnat_of_nat x = whn)";
       
   869 by (auto_tac (claset() addDs [FreeUltrafilterNat_not_finite], 
       
   870               simpset()));
       
   871 qed "not_ex_hypnat_of_nat_eq_omega";
       
   872 
       
   873 Goal "hypnat_of_nat x ~= whn";
       
   874 by (cut_facts_tac [not_ex_hypnat_of_nat_eq_omega] 1);
       
   875 by (Auto_tac);
       
   876 qed "hypnat_of_nat_not_eq_omega";
       
   877 Addsimps [hypnat_of_nat_not_eq_omega RS not_sym];
       
   878 
       
   879 (*-----------------------------------------------------------
       
   880     Properties of the set SHNat of embedded natural numbers
       
   881               (cf. set SReal in NSA.thy/NSA.ML)
       
   882  ----------------------------------------------------------*)
       
   883 
       
   884 (* Infinite hypernatural not in embedded SHNat *)
       
   885 Goalw [SHNat_def] "whn ~: SHNat";
       
   886 by (Auto_tac);
       
   887 qed "SHNAT_omega_not_mem";
       
   888 Addsimps [SHNAT_omega_not_mem];
       
   889 
       
   890 (*-----------------------------------------------------------------------
       
   891      Closure laws for members of (embedded) set standard naturals SHNat
       
   892  -----------------------------------------------------------------------*)
       
   893 Goalw [SHNat_def] 
       
   894       "!!x. [| x: SHNat; y: SHNat |] ==> x + y: SHNat";
       
   895 by (Step_tac 1);
       
   896 by (res_inst_tac [("x","N + Na")] exI 1);
       
   897 by (simp_tac (simpset() addsimps [hypnat_of_nat_add]) 1);
       
   898 qed "SHNat_add";
       
   899 
       
   900 Goalw [SHNat_def] 
       
   901       "!!x. [| x: SHNat; y: SHNat |] ==> x - y: SHNat";
       
   902 by (Step_tac 1);
       
   903 by (res_inst_tac [("x","N - Na")] exI 1);
       
   904 by (simp_tac (simpset() addsimps [hypnat_of_nat_minus]) 1);
       
   905 qed "SHNat_minus";
       
   906 
       
   907 Goalw [SHNat_def] 
       
   908       "!!x. [| x: SHNat; y: SHNat |] ==> x * y: SHNat";
       
   909 by (Step_tac 1);
       
   910 by (res_inst_tac [("x","N * Na")] exI 1);
       
   911 by (simp_tac (simpset() addsimps [hypnat_of_nat_mult]) 1);
       
   912 qed "SHNat_mult";
       
   913 
       
   914 Goal "!!x. [| x + y : SHNat; y: SHNat |] ==> x: SHNat";
       
   915 by (dres_inst_tac [("x","x+y")] SHNat_minus 1);
       
   916 by (Auto_tac);
       
   917 qed "SHNat_add_cancel";
       
   918 
       
   919 Goalw [SHNat_def] "hypnat_of_nat x : SHNat";
       
   920 by (Blast_tac 1);
       
   921 qed "SHNat_hypnat_of_nat";
       
   922 Addsimps [SHNat_hypnat_of_nat];
       
   923 
       
   924 Goal "hypnat_of_nat 1 : SHNat";
       
   925 by (Simp_tac 1);
       
   926 qed "SHNat_hypnat_of_nat_one";
       
   927 
       
   928 Goal "hypnat_of_nat 0 : SHNat";
       
   929 by (Simp_tac 1);
       
   930 qed "SHNat_hypnat_of_nat_zero";
       
   931 
       
   932 Goal "1hn : SHNat";
       
   933 by (simp_tac (simpset() addsimps [SHNat_hypnat_of_nat_one,
       
   934     hypnat_of_nat_one RS sym]) 1);
       
   935 qed "SHNat_one";
       
   936 
       
   937 Goal "0 : SHNat";
       
   938 by (simp_tac (simpset() addsimps [SHNat_hypnat_of_nat_zero,
       
   939     hypnat_of_nat_zero RS sym]) 1);
       
   940 qed "SHNat_zero";
       
   941 
       
   942 Addsimps [SHNat_hypnat_of_nat_one,SHNat_hypnat_of_nat_zero,
       
   943           SHNat_one,SHNat_zero];
       
   944 
       
   945 Goal "1hn + 1hn : SHNat";
       
   946 by (rtac ([SHNat_one,SHNat_one] MRS SHNat_add) 1);
       
   947 qed "SHNat_two";
       
   948 
       
   949 Goalw [SHNat_def] "{x. hypnat_of_nat x : SHNat} = (UNIV::nat set)";
       
   950 by (Auto_tac);
       
   951 qed "SHNat_UNIV_nat";
       
   952 
       
   953 Goalw [SHNat_def] "(x: SHNat) = (EX y. x = hypnat_of_nat  y)";
       
   954 by (Auto_tac);
       
   955 qed "SHNat_iff";
       
   956 
       
   957 Goalw [SHNat_def] "hypnat_of_nat ``(UNIV::nat set) = SHNat";
       
   958 by (Auto_tac);
       
   959 qed "hypnat_of_nat_image";
       
   960 
       
   961 Goalw [SHNat_def] "inv hypnat_of_nat ``SHNat = (UNIV::nat set)";
       
   962 by (Auto_tac);
       
   963 by (rtac (inj_hypnat_of_nat RS inv_f_f RS subst) 1);
       
   964 by (Blast_tac 1);
       
   965 qed "inv_hypnat_of_nat_image";
       
   966 
       
   967 Goalw [SHNat_def] 
       
   968       "!!P. [| EX x. x: P; P <= SHNat |] ==> \
       
   969 \           EX Q. P = hypnat_of_nat `` Q";
       
   970 by (Best_tac 1); 
       
   971 qed "SHNat_hypnat_of_nat_image";
       
   972 
       
   973 Goalw [SHNat_def] 
       
   974       "SHNat = hypnat_of_nat `` (UNIV::nat set)";
       
   975 by (Auto_tac);
       
   976 qed "SHNat_hypnat_of_nat_iff";
       
   977 
       
   978 Goalw [SHNat_def] "SHNat <= (UNIV::hypnat set)";
       
   979 by (Auto_tac);
       
   980 qed "SHNat_subset_UNIV";
       
   981 
       
   982 Goal "{n. n <= Suc m} = {n. n <= m} Un {n. n = Suc m}";
       
   983 by (auto_tac (claset(),simpset() addsimps [le_Suc_eq]));
       
   984 qed "leSuc_Un_eq";
       
   985 
       
   986 Goal "finite {n::nat. n <= m}";
       
   987 by (nat_ind_tac "m" 1);
       
   988 by (auto_tac (claset(),simpset() addsimps [leSuc_Un_eq]));
       
   989 qed "finite_nat_le_segment";
       
   990 
       
   991 Goal "{n::nat. m < n} : FreeUltrafilterNat";
       
   992 by (cut_inst_tac [("m2","m")] (finite_nat_le_segment RS 
       
   993     FreeUltrafilterNat_finite RS FreeUltrafilterNat_Compl_mem) 1);
       
   994 by (Fuf_tac 1);
       
   995 qed "lemma_unbounded_set";
       
   996 Addsimps [lemma_unbounded_set];
       
   997 
       
   998 Goalw [SHNat_def,hypnat_of_nat_def,
       
   999            hypnat_less_def,hypnat_omega_def] 
       
  1000            "ALL n: SHNat. n < whn";
       
  1001 by (Clarify_tac 1);
       
  1002 by (auto_tac (claset() addSIs [exI],simpset()));
       
  1003 qed "hypnat_omega_gt_SHNat";
       
  1004 
       
  1005 Goal "hypnat_of_nat n < whn";
       
  1006 by (cut_facts_tac [hypnat_omega_gt_SHNat] 1);
       
  1007 by (dres_inst_tac [("x","hypnat_of_nat n")] bspec 1);
       
  1008 by (Auto_tac);
       
  1009 qed "hypnat_of_nat_less_whn";
       
  1010 Addsimps [hypnat_of_nat_less_whn];
       
  1011 
       
  1012 Goal "hypnat_of_nat n <= whn";
       
  1013 by (rtac (hypnat_of_nat_less_whn RS hypnat_less_imp_le) 1);
       
  1014 qed "hypnat_of_nat_le_whn";
       
  1015 Addsimps [hypnat_of_nat_le_whn];
       
  1016 
       
  1017 Goal "0 < whn";
       
  1018 by (rtac (hypnat_omega_gt_SHNat RS ballE) 1);
       
  1019 by (Auto_tac);
       
  1020 qed "hypnat_zero_less_hypnat_omega";
       
  1021 Addsimps [hypnat_zero_less_hypnat_omega];
       
  1022 
       
  1023 Goal "1hn < whn";
       
  1024 by (rtac (hypnat_omega_gt_SHNat RS ballE) 1);
       
  1025 by (Auto_tac);
       
  1026 qed "hypnat_one_less_hypnat_omega";
       
  1027 Addsimps [hypnat_one_less_hypnat_omega];
       
  1028 
       
  1029 (*--------------------------------------------------------------------------
       
  1030      Theorems about infinite hypernatural numbers -- HNatInfinite
       
  1031  -------------------------------------------------------------------------*)
       
  1032 Goalw [HNatInfinite_def,SHNat_def] "whn : HNatInfinite";
       
  1033 by (Auto_tac);
       
  1034 qed "HNatInfinite_whn";
       
  1035 Addsimps [HNatInfinite_whn];
       
  1036 
       
  1037 Goalw [HNatInfinite_def] "!!x. x: SHNat ==> x ~: HNatInfinite";
       
  1038 by (Simp_tac 1);
       
  1039 qed "SHNat_not_HNatInfinite";
       
  1040 
       
  1041 Goalw [HNatInfinite_def] "!!x. x ~: HNatInfinite ==> x: SHNat";
       
  1042 by (Asm_full_simp_tac 1);
       
  1043 qed "not_HNatInfinite_SHNat";
       
  1044 
       
  1045 Goalw [HNatInfinite_def] "!!x. x ~: SHNat ==> x: HNatInfinite";
       
  1046 by (Simp_tac 1);
       
  1047 qed "not_SHNat_HNatInfinite";
       
  1048 
       
  1049 Goalw [HNatInfinite_def] "!!x. x: HNatInfinite ==> x ~: SHNat";
       
  1050 by (Asm_full_simp_tac 1);
       
  1051 qed "HNatInfinite_not_SHNat";
       
  1052 
       
  1053 Goal "(x: SHNat) = (x ~: HNatInfinite)";
       
  1054 by (blast_tac (claset() addSIs [SHNat_not_HNatInfinite,
       
  1055     not_HNatInfinite_SHNat]) 1);
       
  1056 qed "SHNat_not_HNatInfinite_iff";
       
  1057 
       
  1058 Goal "(x ~: SHNat) = (x: HNatInfinite)";
       
  1059 by (blast_tac (claset() addSIs [not_SHNat_HNatInfinite,
       
  1060     HNatInfinite_not_SHNat]) 1);
       
  1061 qed "not_SHNat_HNatInfinite_iff";
       
  1062 
       
  1063 Goal "x : SHNat | x : HNatInfinite";
       
  1064 by (simp_tac (simpset() addsimps [SHNat_not_HNatInfinite_iff]) 1);
       
  1065 qed "SHNat_HNatInfinite_disj";
       
  1066 
       
  1067 (*-------------------------------------------------------------------
       
  1068    Proof of alternative definition for set of Infinite hypernatural 
       
  1069    numbers --- HNatInfinite = {N. ALL n: SHNat. n < N}
       
  1070  -------------------------------------------------------------------*)
       
  1071 Goal "!!N (xa::nat=>nat). \
       
  1072 \         (ALL N. {n. xa n ~= N} : FreeUltrafilterNat) ==> \
       
  1073 \         ({n. N < xa n} : FreeUltrafilterNat)";
       
  1074 by (nat_ind_tac "N" 1);
       
  1075 by (dres_inst_tac [("x","0")] spec 1);
       
  1076 by (rtac ccontr 1 THEN dtac FreeUltrafilterNat_Compl_mem 1
       
  1077     THEN dtac FreeUltrafilterNat_Int 1 THEN assume_tac 1);
       
  1078 by (Asm_full_simp_tac 1);
       
  1079 by (dres_inst_tac [("x","Suc N")] spec 1);
       
  1080 by (fuf_tac (claset() addSDs [Suc_leI],simpset() addsimps 
       
  1081     [le_eq_less_or_eq]) 1);
       
  1082 qed "HNatInfinite_FreeUltrafilterNat_lemma";
       
  1083 
       
  1084 (*** alternative definition ***)
       
  1085 Goalw [HNatInfinite_def,SHNat_def,hypnat_of_nat_def] 
       
  1086       "HNatInfinite = {N. ALL n:SHNat. n < N}";
       
  1087 by (Step_tac 1);
       
  1088 by (dres_inst_tac [("x","Abs_hypnat \
       
  1089 \        (hypnatrel ^^ {%n. N})")] bspec 2);
       
  1090 by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
       
  1091 by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
       
  1092 by (auto_tac (claset(),simpset() addsimps [hypnat_less_iff]));
       
  1093 by (auto_tac (claset() addSIs [exI] addEs 
       
  1094     [HNatInfinite_FreeUltrafilterNat_lemma],
       
  1095     simpset() addsimps [FreeUltrafilterNat_Compl_iff1, 
       
  1096      CLAIM "- {n. xa n = N} = {n. xa n ~= N}"]));
       
  1097 qed "HNatInfinite_iff";
       
  1098 
       
  1099 (*--------------------------------------------------------------------
       
  1100    Alternative definition for HNatInfinite using Free ultrafilter
       
  1101  --------------------------------------------------------------------*)
       
  1102 Goal "!!x. x : HNatInfinite ==> EX X: Rep_hypnat x. \
       
  1103 \          ALL u. {n. u < X n}:  FreeUltrafilterNat";
       
  1104 by (auto_tac (claset(),simpset() addsimps [hypnat_less_def,
       
  1105     HNatInfinite_iff,SHNat_iff,hypnat_of_nat_def]));
       
  1106 by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
       
  1107 by (EVERY[Auto_tac, rtac bexI 1, 
       
  1108     rtac lemma_hypnatrel_refl 2, Step_tac 1]);
       
  1109 by (dres_inst_tac [("x","hypnat_of_nat u")] bspec 1);
       
  1110 by (Simp_tac 1);
       
  1111 by (auto_tac (claset(),
       
  1112     simpset() addsimps [hypnat_of_nat_def]));
       
  1113 by (Fuf_tac 1);
       
  1114 qed "HNatInfinite_FreeUltrafilterNat";
       
  1115 
       
  1116 Goal "!!x. EX X: Rep_hypnat x. \
       
  1117 \          ALL u. {n. u < X n}:  FreeUltrafilterNat \
       
  1118 \          ==> x: HNatInfinite";
       
  1119 by (auto_tac (claset(),simpset() addsimps [hypnat_less_def,
       
  1120     HNatInfinite_iff,SHNat_iff,hypnat_of_nat_def]));
       
  1121 by (rtac exI 1 THEN Auto_tac);
       
  1122 qed "FreeUltrafilterNat_HNatInfinite";
       
  1123 
       
  1124 Goal "!!x. (x : HNatInfinite) = (EX X: Rep_hypnat x. \
       
  1125 \          ALL u. {n. u < X n}:  FreeUltrafilterNat)";
       
  1126 by (blast_tac (claset() addIs [HNatInfinite_FreeUltrafilterNat,
       
  1127     FreeUltrafilterNat_HNatInfinite]) 1);
       
  1128 qed "HNatInfinite_FreeUltrafilterNat_iff";
       
  1129 
       
  1130 Goal "!!x. x : HNatInfinite ==> 1hn < x";
       
  1131 by (auto_tac (claset(),simpset() addsimps [HNatInfinite_iff]));
       
  1132 qed "HNatInfinite_gt_one";
       
  1133 Addsimps [HNatInfinite_gt_one];
       
  1134 
       
  1135 Goal "!!x. 0 ~: HNatInfinite";
       
  1136 by (auto_tac (claset(),simpset() 
       
  1137     addsimps [HNatInfinite_iff]));
       
  1138 by (dres_inst_tac [("a","1hn")] equals0D 1);
       
  1139 by (Asm_full_simp_tac 1);
       
  1140 qed "zero_not_mem_HNatInfinite";
       
  1141 Addsimps [zero_not_mem_HNatInfinite];
       
  1142 
       
  1143 Goal "!!x. x : HNatInfinite ==> x ~= 0";
       
  1144 by (Auto_tac);
       
  1145 qed "HNatInfinite_not_eq_zero";
       
  1146 
       
  1147 Goal "!!x. x : HNatInfinite ==> 1hn <= x";
       
  1148 by (blast_tac (claset() addIs [hypnat_less_imp_le,
       
  1149          HNatInfinite_gt_one]) 1);
       
  1150 qed "HNatInfinite_ge_one";
       
  1151 Addsimps [HNatInfinite_ge_one];
       
  1152 
       
  1153 (*--------------------------------------------------
       
  1154                    Closure Rules
       
  1155  --------------------------------------------------*)
       
  1156 Goal "!!x. [| x: HNatInfinite; y: HNatInfinite |] \
       
  1157 \           ==> x + y: HNatInfinite";
       
  1158 by (auto_tac (claset(),simpset() addsimps [HNatInfinite_iff]));
       
  1159 by (dtac bspec 1 THEN assume_tac 1);
       
  1160 by (dtac (SHNat_zero RSN (2,bspec)) 1);
       
  1161 by (dtac hypnat_add_less_mono 1 THEN assume_tac 1);
       
  1162 by (Asm_full_simp_tac 1);
       
  1163 qed "HNatInfinite_add";
       
  1164 
       
  1165 Goal "!!x. [| x: HNatInfinite; y: SHNat |] \
       
  1166 \                       ==> x + y: HNatInfinite";
       
  1167 by (rtac ccontr 1 THEN dtac not_HNatInfinite_SHNat 1);
       
  1168 by (dres_inst_tac [("x","x + y")] SHNat_minus 1);
       
  1169 by (auto_tac (claset(),simpset() addsimps 
       
  1170     [SHNat_not_HNatInfinite_iff]));
       
  1171 qed "HNatInfinite_SHNat_add";
       
  1172 
       
  1173 goal HyperNat.thy "!!x. [| x: HNatInfinite; y: SHNat |] \
       
  1174 \                       ==> x - y: HNatInfinite";
       
  1175 by (rtac ccontr 1 THEN dtac not_HNatInfinite_SHNat 1);
       
  1176 by (dres_inst_tac [("x","x - y")] SHNat_add 1);
       
  1177 by (subgoal_tac "y <= x" 2);
       
  1178 by (auto_tac (claset() addSDs [hypnat_le_add_diff_inverse2],
       
  1179     simpset() addsimps [not_SHNat_HNatInfinite_iff RS sym]));
       
  1180 by (auto_tac (claset() addSIs [hypnat_less_imp_le],
       
  1181     simpset() addsimps [not_SHNat_HNatInfinite_iff,HNatInfinite_iff]));
       
  1182 qed "HNatInfinite_SHNat_diff";
       
  1183 
       
  1184 Goal 
       
  1185      "!!x. x: HNatInfinite ==> x + 1hn: HNatInfinite";
       
  1186 by (auto_tac (claset() addIs [HNatInfinite_SHNat_add],
       
  1187               simpset()));
       
  1188 qed "HNatInfinite_add_one";
       
  1189 
       
  1190 Goal 
       
  1191      "!!x. x: HNatInfinite ==> x - 1hn: HNatInfinite";
       
  1192 by (rtac ccontr 1 THEN dtac not_HNatInfinite_SHNat 1);
       
  1193 by (dres_inst_tac [("x","x - 1hn"),("y","1hn")] SHNat_add 1);
       
  1194 by (auto_tac (claset(),simpset() addsimps 
       
  1195     [not_SHNat_HNatInfinite_iff RS sym]));
       
  1196 qed "HNatInfinite_minus_one";
       
  1197 
       
  1198 Goal "!!x. x : HNatInfinite ==> EX y. x = y + 1hn";
       
  1199 by (res_inst_tac [("x","x - 1hn")] exI 1);
       
  1200 by (Auto_tac);
       
  1201 qed "HNatInfinite_is_Suc";
       
  1202 
       
  1203 (*---------------------------------------------------------------
       
  1204        HNat : the hypernaturals embedded in the hyperreals
       
  1205        Obtained using the NS extension of the naturals
       
  1206  --------------------------------------------------------------*)
       
  1207  
       
  1208 Goalw [HNat_def,starset_def,
       
  1209          hypreal_of_posnat_def,hypreal_of_real_def] 
       
  1210          "hypreal_of_posnat N : HNat";
       
  1211 by (Auto_tac);
       
  1212 by (Ultra_tac 1);
       
  1213 by (res_inst_tac [("x","Suc N")] exI 1);
       
  1214 by (dtac sym 1 THEN auto_tac (claset(),simpset() 
       
  1215     addsimps [real_of_preal_real_of_nat_Suc]));
       
  1216 qed "HNat_hypreal_of_posnat";
       
  1217 Addsimps [HNat_hypreal_of_posnat];
       
  1218 
       
  1219 Goalw [HNat_def,starset_def]
       
  1220      "[| x: HNat; y: HNat |] ==> x + y: HNat";
       
  1221 by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
       
  1222 by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
       
  1223 by (auto_tac (claset() addSDs [bspec] addIs [lemma_hyprel_refl],
       
  1224     simpset() addsimps [hypreal_add]));
       
  1225 by (Ultra_tac 1);
       
  1226 by (dres_inst_tac [("t","Y xb")] sym 1);
       
  1227 by (auto_tac (claset(),simpset() addsimps [real_of_nat_add RS sym]));
       
  1228 qed "HNat_add";
       
  1229 
       
  1230 Goalw [HNat_def,starset_def]
       
  1231      "[| x: HNat; y: HNat |] ==> x * y: HNat";
       
  1232 by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
       
  1233 by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
       
  1234 by (auto_tac (claset() addSDs [bspec] addIs [lemma_hyprel_refl],
       
  1235     simpset() addsimps [hypreal_mult]));
       
  1236 by (Ultra_tac 1);
       
  1237 by (dres_inst_tac [("t","Y xb")] sym 1);
       
  1238 by (auto_tac (claset(),simpset() addsimps [real_of_nat_mult RS sym]));
       
  1239 qed "HNat_mult";
       
  1240 
       
  1241 (*---------------------------------------------------------------
       
  1242     Embedding of the hypernaturals into the hyperreal
       
  1243  --------------------------------------------------------------*)
       
  1244 (*** A lemma that should have been derived a long time ago! ***)
       
  1245 Goal "(Ya : hyprel ^^{%n. f(n)}) = \
       
  1246 \         ({n. f n = Ya n} : FreeUltrafilterNat)";
       
  1247 by (Auto_tac);
       
  1248 qed "lemma_hyprel_FUFN";
       
  1249 
       
  1250 Goalw [hypreal_of_hypnat_def]
       
  1251       "hypreal_of_hypnat (Abs_hypnat(hypnatrel^^{%n. X n})) = \
       
  1252 \         Abs_hypreal(hyprel ^^ {%n. real_of_nat (X n)})";
       
  1253 by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1);
       
  1254 by (auto_tac (claset() addEs [FreeUltrafilterNat_Int RS
       
  1255     FreeUltrafilterNat_subset],simpset() addsimps 
       
  1256     [lemma_hyprel_FUFN]));
       
  1257 qed "hypreal_of_hypnat";
       
  1258 
       
  1259 Goal "inj(hypreal_of_hypnat)";
       
  1260 by (rtac injI 1);
       
  1261 by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
       
  1262 by (res_inst_tac [("z","y")] eq_Abs_hypnat 1);
       
  1263 by (auto_tac (claset(),simpset() addsimps 
       
  1264     [hypreal_of_hypnat,real_of_nat_eq_cancel]));
       
  1265 qed "inj_hypreal_of_hypnat";
       
  1266 
       
  1267 Goal "(hypreal_of_hypnat n = hypreal_of_hypnat m) = (n = m)";
       
  1268 by (auto_tac (claset(),simpset() addsimps [inj_hypreal_of_hypnat RS injD]));
       
  1269 qed "hypreal_of_hypnat_eq_cancel";
       
  1270 Addsimps [hypreal_of_hypnat_eq_cancel];
       
  1271 
       
  1272 Goal "(hypnat_of_nat n = hypnat_of_nat m) = (n = m)";
       
  1273 by (auto_tac (claset() addDs [inj_hypnat_of_nat RS injD],
       
  1274               simpset()));
       
  1275 qed "hypnat_of_nat_eq_cancel";
       
  1276 Addsimps [hypnat_of_nat_eq_cancel];
       
  1277 
       
  1278 Goalw [hypreal_zero_def,hypnat_zero_def] 
       
  1279            "hypreal_of_hypnat  0 = 0";
       
  1280 by (simp_tac (simpset() addsimps [hypreal_of_hypnat,
       
  1281     real_of_nat_zero]) 1);
       
  1282 qed "hypreal_of_hypnat_zero";
       
  1283 
       
  1284 Goalw [hypreal_one_def,hypnat_one_def] 
       
  1285            "hypreal_of_hypnat  1hn = 1hr";
       
  1286 by (simp_tac (simpset() addsimps [hypreal_of_hypnat,
       
  1287     real_of_nat_one]) 1);
       
  1288 qed "hypreal_of_hypnat_one";
       
  1289 
       
  1290 Goal "hypreal_of_hypnat n1 + hypreal_of_hypnat n2 = hypreal_of_hypnat (n1 + n2)";
       
  1291 by (res_inst_tac [("z","n1")] eq_Abs_hypnat 1);
       
  1292 by (res_inst_tac [("z","n2")] eq_Abs_hypnat 1);
       
  1293 by (asm_simp_tac (simpset() addsimps [hypreal_of_hypnat,
       
  1294         hypreal_add,hypnat_add,real_of_nat_add]) 1);
       
  1295 qed "hypreal_of_hypnat_add";
       
  1296 
       
  1297 Goal "hypreal_of_hypnat n1 * hypreal_of_hypnat n2 = hypreal_of_hypnat (n1 * n2)";
       
  1298 by (res_inst_tac [("z","n1")] eq_Abs_hypnat 1);
       
  1299 by (res_inst_tac [("z","n2")] eq_Abs_hypnat 1);
       
  1300 by (asm_simp_tac (simpset() addsimps [hypreal_of_hypnat,
       
  1301         hypreal_mult,hypnat_mult,real_of_nat_mult]) 1);
       
  1302 qed "hypreal_of_hypnat_mult";
       
  1303 
       
  1304 Goal "(hypreal_of_hypnat n < hypreal_of_hypnat m) = (n < m)";
       
  1305 by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
       
  1306 by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
       
  1307 by (asm_simp_tac (simpset() addsimps 
       
  1308     [hypreal_of_hypnat,hypreal_less,hypnat_less]) 1);
       
  1309 qed "hypreal_of_hypnat_less_iff";
       
  1310 Addsimps [hypreal_of_hypnat_less_iff];
       
  1311 
       
  1312 Goal "(hypreal_of_hypnat N = 0) = (N = 0)";
       
  1313 by (simp_tac (simpset() addsimps [hypreal_of_hypnat_zero RS sym]) 1);
       
  1314 qed "hypreal_of_hypnat_eq_zero_iff";
       
  1315 Addsimps [hypreal_of_hypnat_eq_zero_iff];
       
  1316 
       
  1317 Goal "ALL n. N <= n ==> N = (0::hypnat)";
       
  1318 by (dres_inst_tac [("x","0")] spec 1);
       
  1319 by (res_inst_tac [("z","N")] eq_Abs_hypnat 1);
       
  1320 by (auto_tac (claset(),simpset() addsimps [hypnat_le,hypnat_zero_def]));
       
  1321 qed "hypnat_eq_zero";
       
  1322 Addsimps [hypnat_eq_zero];
       
  1323 
       
  1324 Goal "~ (ALL n. n = (0::hypnat))";
       
  1325 by Auto_tac;
       
  1326 by (res_inst_tac [("x","1hn")] exI 1);
       
  1327 by (Simp_tac 1);
       
  1328 qed "hypnat_not_all_eq_zero";
       
  1329 Addsimps [hypnat_not_all_eq_zero];
       
  1330 
       
  1331 Goal "n ~= 0 ==> (n <= 1hn) = (n = 1hn)";
       
  1332 by (auto_tac (claset(),simpset() addsimps [hypnat_le_eq_less_or_eq]));
       
  1333 qed "hypnat_le_one_eq_one";
       
  1334 Addsimps [hypnat_le_one_eq_one];
       
  1335 
       
  1336 
       
  1337 
       
  1338