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