src/HOL/Real/PRat.ML
author paulson
Fri, 02 Nov 2001 17:55:24 +0100
changeset 12018 ec054019c910
parent 11701 3d51fbf81c17
child 13438 527811f00c56
permissions -rw-r--r--
Numerals and simprocs for types real and hypreal. The abstract constants 0, 1 and binary numerals work harmoniously.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     1
(*  Title       : PRat.ML
7219
4e3f386c2e37 inserted Id: lines
paulson
parents: 7077
diff changeset
     2
    ID          : $Id$
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     3
    Author      : Jacques D. Fleuriot
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     4
    Copyright   : 1998  University of Cambridge
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     5
    Description : The positive rationals
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     6
*) 
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     7
9426
d29faa6cc391 avoid referencing thy value;
wenzelm
parents: 9391
diff changeset
     8
(*** Many theorems similar to those in theory Integ ***)
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     9
(*** Proving that ratrel is an equivalence relation ***)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    10
7219
4e3f386c2e37 inserted Id: lines
paulson
parents: 7077
diff changeset
    11
Goal "[| (x1::pnat) * y2 = x2 * y1; x2 * y3 = x3 * y2 |] \
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    12
\            ==> x1 * y3 = x3 * y1";        
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    13
by (res_inst_tac [("k1","y2")] (pnat_mult_cancel1 RS iffD1) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    14
by (auto_tac (claset(), simpset() addsimps [pnat_mult_assoc RS sym]));
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    15
by (auto_tac (claset(),simpset() addsimps [pnat_mult_commute]));
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    16
by (dres_inst_tac [("s","x2 * y3")] sym 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    17
by (asm_simp_tac (simpset() addsimps [pnat_mult_left_commute,
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    18
    pnat_mult_commute]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    19
qed "prat_trans_lemma";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    20
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    21
(** Natural deduction for ratrel **)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    22
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    23
Goalw [ratrel_def]
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    24
    "(((x1,y1),(x2,y2)): ratrel) = (x1 * y2 = x2 * y1)";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    25
by (Fast_tac 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    26
qed "ratrel_iff";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    27
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    28
Goalw [ratrel_def]
5148
74919e8f221c More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5143
diff changeset
    29
    "[| x1 * y2 = x2 * y1 |] ==> ((x1,y1),(x2,y2)): ratrel";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    30
by (Fast_tac  1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    31
qed "ratrelI";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    32
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    33
Goalw [ratrel_def]
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    34
  "p: ratrel --> (EX x1 y1 x2 y2. \
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    35
\                  p = ((x1,y1),(x2,y2)) & x1 *y2 = x2 *y1)";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    36
by (Fast_tac 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    37
qed "ratrelE_lemma";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    38
9969
4753185f1dd2 renamed (most of...) the select rules
paulson
parents: 9426
diff changeset
    39
val [major,minor] = Goal
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    40
  "[| p: ratrel;  \
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    41
\     !!x1 y1 x2 y2. [| p = ((x1,y1),(x2,y2));  x1*y2 = x2*y1 \
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    42
\                    |] ==> Q |] ==> Q";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    43
by (cut_facts_tac [major RS (ratrelE_lemma RS mp)] 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    44
by (REPEAT (eresolve_tac [asm_rl,exE,conjE,minor] 1));
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    45
qed "ratrelE";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    46
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    47
AddSIs [ratrelI];
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    48
AddSEs [ratrelE];
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    49
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    50
Goal "(x,x): ratrel";
9369
139fde7af7bd use pair_tac;
wenzelm
parents: 9108
diff changeset
    51
by (pair_tac "x" 1);
139fde7af7bd use pair_tac;
wenzelm
parents: 9108
diff changeset
    52
by (rtac ratrelI 1);
139fde7af7bd use pair_tac;
wenzelm
parents: 9108
diff changeset
    53
by (rtac refl 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    54
qed "ratrel_refl";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    55
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    56
Goalw [equiv_def, refl_def, sym_def, trans_def]
7376
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
    57
    "equiv UNIV ratrel";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    58
by (fast_tac (claset() addSIs [ratrel_refl] 
7376
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
    59
                       addSEs [sym, prat_trans_lemma]) 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    60
qed "equiv_ratrel";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    61
9108
9fff97d29837 bind_thm(s);
wenzelm
parents: 9043
diff changeset
    62
bind_thm ("equiv_ratrel_iff", [equiv_ratrel, UNIV_I, UNIV_I] MRS eq_equiv_class_iff);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    63
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
    64
Goalw  [prat_def,ratrel_def,quotient_def] "ratrel``{(x,y)}:prat";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    65
by (Blast_tac 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    66
qed "ratrel_in_prat";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    67
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    68
Goal "inj_on Abs_prat prat";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    69
by (rtac inj_on_inverseI 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    70
by (etac Abs_prat_inverse 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    71
qed "inj_on_Abs_prat";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    72
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    73
Addsimps [equiv_ratrel_iff,inj_on_Abs_prat RS inj_on_iff,
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    74
          ratrel_iff, ratrel_in_prat, Abs_prat_inverse];
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    75
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    76
Addsimps [equiv_ratrel RS eq_equiv_class_iff];
9108
9fff97d29837 bind_thm(s);
wenzelm
parents: 9043
diff changeset
    77
bind_thm ("eq_ratrelD", equiv_ratrel RSN (2,eq_equiv_class));
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    78
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    79
Goal "inj(Rep_prat)";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    80
by (rtac inj_inverseI 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    81
by (rtac Rep_prat_inverse 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    82
qed "inj_Rep_prat";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    83
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
    84
(** prat_of_pnat: the injection from pnat to prat **)
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
    85
Goal "inj(prat_of_pnat)";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    86
by (rtac injI 1);
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
    87
by (rewtac prat_of_pnat_def);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    88
by (dtac (inj_on_Abs_prat RS inj_onD) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    89
by (REPEAT (rtac ratrel_in_prat 1));
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    90
by (dtac eq_equiv_class 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    91
by (rtac equiv_ratrel 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    92
by (Fast_tac 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    93
by Safe_tac;
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    94
by (Asm_full_simp_tac 1);
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
    95
qed "inj_prat_of_pnat";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    96
9969
4753185f1dd2 renamed (most of...) the select rules
paulson
parents: 9426
diff changeset
    97
val [prem] = Goal
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
    98
    "(!!x y. z = Abs_prat(ratrel``{(x,y)}) ==> P) ==> P";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    99
by (res_inst_tac [("x1","z")] 
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   100
    (rewrite_rule [prat_def] Rep_prat RS quotientE) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   101
by (dres_inst_tac [("f","Abs_prat")] arg_cong 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   102
by (res_inst_tac [("p","x")] PairE 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   103
by (rtac prem 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   104
by (asm_full_simp_tac (simpset() addsimps [Rep_prat_inverse]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   105
qed "eq_Abs_prat";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   106
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   107
(**** qinv: inverse on prat ****)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   108
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
   109
Goalw [congruent_def] "congruent ratrel (%(x,y). ratrel``{(y,x)})";
10232
529c65b5dcde restoration of "equalityI"; renaming of contrapos rules
paulson
parents: 9969
diff changeset
   110
by (auto_tac (claset(), simpset() addsimps [pnat_mult_commute]));  
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   111
qed "qinv_congruent";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   112
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   113
Goalw [qinv_def]
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
   114
      "qinv (Abs_prat(ratrel``{(x,y)})) = Abs_prat(ratrel `` {(y,x)})";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   115
by (simp_tac (simpset() addsimps 
7376
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   116
	      [equiv_ratrel RS UN_equiv_class, qinv_congruent]) 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   117
qed "qinv";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   118
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   119
Goal "qinv (qinv z) = z";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   120
by (res_inst_tac [("z","z")] eq_Abs_prat 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   121
by (asm_simp_tac (simpset() addsimps [qinv]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   122
qed "qinv_qinv";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   123
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   124
Goal "inj(qinv)";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   125
by (rtac injI 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   126
by (dres_inst_tac [("f","qinv")] arg_cong 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   127
by (asm_full_simp_tac (simpset() addsimps [qinv_qinv]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   128
qed "inj_qinv";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   129
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
   130
Goalw [prat_of_pnat_def] 
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11655
diff changeset
   131
      "qinv(prat_of_pnat  (Abs_pnat (Suc 0))) = prat_of_pnat (Abs_pnat (Suc 0))";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   132
by (simp_tac (simpset() addsimps [qinv]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   133
qed "qinv_1";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   134
5148
74919e8f221c More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5143
diff changeset
   135
Goal "!!(x1::pnat). [| x1 * y2 = x2 * y1 |] ==> \
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   136
\     (x * y1 + x1 * ya) * (ya * y2) = (x * y2 + x2 * ya) * (ya * y1)";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   137
by (auto_tac (claset() addSIs [pnat_same_multI2],
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   138
       simpset() addsimps [pnat_add_mult_distrib,
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   139
       pnat_mult_assoc]));
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   140
by (res_inst_tac [("n1","y2")] (pnat_mult_commute RS subst) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   141
by (auto_tac (claset() addIs [pnat_add_left_cancel RS iffD2],simpset() addsimps pnat_mult_ac));
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   142
by (res_inst_tac [("y1","x1")] (pnat_mult_left_commute RS subst) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   143
by (res_inst_tac [("y1","x1")] (pnat_mult_left_commute RS ssubst) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   144
by (auto_tac (claset(),simpset() addsimps [pnat_mult_assoc RS sym]));
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   145
qed "prat_add_congruent2_lemma";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   146
7219
4e3f386c2e37 inserted Id: lines
paulson
parents: 7077
diff changeset
   147
Goal "congruent2 ratrel (%p1 p2.                  \
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
   148
\        (%(x1,y1). (%(x2,y2). ratrel``{(x1*y2 + x2*y1, y1*y2)}) p2) p1)";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   149
by (rtac (equiv_ratrel RS congruent2_commuteI) 1);
10232
529c65b5dcde restoration of "equalityI"; renaming of contrapos rules
paulson
parents: 9969
diff changeset
   150
by (auto_tac (claset() delrules [equalityI],
529c65b5dcde restoration of "equalityI"; renaming of contrapos rules
paulson
parents: 9969
diff changeset
   151
              simpset() addsimps [prat_add_congruent2_lemma]));
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   152
by (asm_simp_tac (simpset() addsimps [pnat_mult_commute,pnat_add_commute]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   153
qed "prat_add_congruent2";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   154
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   155
Goalw [prat_add_def]
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
   156
   "Abs_prat((ratrel``{(x1,y1)})) + Abs_prat((ratrel``{(x2,y2)})) =   \
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
   157
\   Abs_prat(ratrel `` {(x1*y2 + x2*y1, y1*y2)})";
7376
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   158
by (simp_tac (simpset() addsimps [UN_UN_split_split_eq, prat_add_congruent2, 
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   159
				  equiv_ratrel RS UN_equiv_class2]) 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   160
qed "prat_add";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   161
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   162
Goal "(z::prat) + w = w + z";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   163
by (res_inst_tac [("z","z")] eq_Abs_prat 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   164
by (res_inst_tac [("z","w")] eq_Abs_prat 1);
5535
678999604ee9 deleted needless parentheses
paulson
parents: 5459
diff changeset
   165
by (asm_simp_tac
678999604ee9 deleted needless parentheses
paulson
parents: 5459
diff changeset
   166
    (simpset() addsimps [prat_add] @ pnat_add_ac @ pnat_mult_ac) 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   167
qed "prat_add_commute";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   168
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   169
Goal "((z1::prat) + z2) + z3 = z1 + (z2 + z3)";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   170
by (res_inst_tac [("z","z1")] eq_Abs_prat 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   171
by (res_inst_tac [("z","z2")] eq_Abs_prat 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   172
by (res_inst_tac [("z","z3")] eq_Abs_prat 1);
5535
678999604ee9 deleted needless parentheses
paulson
parents: 5459
diff changeset
   173
by (asm_simp_tac (simpset() addsimps [pnat_add_mult_distrib2,prat_add] @ 
7376
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   174
		                     pnat_add_ac @ pnat_mult_ac) 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   175
qed "prat_add_assoc";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   176
7376
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   177
(*For AC rewriting*)
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   178
Goal "(z1::prat) + (z2 + z3) = z2 + (z1 + z3)";
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   179
by (rtac (prat_add_commute RS trans) 1);
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   180
by (rtac (prat_add_assoc RS trans) 1);
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   181
by (rtac (prat_add_commute RS arg_cong) 1);
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   182
qed "prat_add_left_commute";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   183
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   184
(* Positive Rational addition is an AC operator *)
9108
9fff97d29837 bind_thm(s);
wenzelm
parents: 9043
diff changeset
   185
bind_thms ("prat_add_ac", [prat_add_assoc, prat_add_commute, prat_add_left_commute]);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   186
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   187
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   188
(*** Congruence property for multiplication ***)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   189
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   190
Goalw [congruent2_def]
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   191
    "congruent2 ratrel (%p1 p2.                  \
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
   192
\         (%(x1,y1). (%(x2,y2). ratrel``{(x1*x2, y1*y2)}) p2) p1)";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   193
(*Proof via congruent2_commuteI seems longer*)
10232
529c65b5dcde restoration of "equalityI"; renaming of contrapos rules
paulson
parents: 9969
diff changeset
   194
by (Clarify_tac 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   195
by (asm_simp_tac (simpset() addsimps [pnat_mult_assoc]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   196
(*The rest should be trivial, but rearranging terms is hard*)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   197
by (res_inst_tac [("x1","x1a")] (pnat_mult_left_commute RS ssubst) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   198
by (asm_simp_tac (simpset() addsimps [pnat_mult_assoc RS sym]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   199
by (asm_simp_tac (simpset() addsimps pnat_mult_ac) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   200
qed "pnat_mult_congruent2";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   201
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   202
Goalw [prat_mult_def]
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
   203
  "Abs_prat(ratrel``{(x1,y1)}) * Abs_prat(ratrel``{(x2,y2)}) = \
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
   204
\  Abs_prat(ratrel``{(x1*x2, y1*y2)})";
7376
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   205
by (asm_simp_tac 
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   206
    (simpset() addsimps [UN_UN_split_split_eq, pnat_mult_congruent2,
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   207
			 equiv_ratrel RS UN_equiv_class2]) 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   208
qed "prat_mult";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   209
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   210
Goal "(z::prat) * w = w * z";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   211
by (res_inst_tac [("z","z")] eq_Abs_prat 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   212
by (res_inst_tac [("z","w")] eq_Abs_prat 1);
5535
678999604ee9 deleted needless parentheses
paulson
parents: 5459
diff changeset
   213
by (asm_simp_tac (simpset() addsimps pnat_mult_ac @ [prat_mult]) 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   214
qed "prat_mult_commute";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   215
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   216
Goal "((z1::prat) * z2) * z3 = z1 * (z2 * z3)";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   217
by (res_inst_tac [("z","z1")] eq_Abs_prat 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   218
by (res_inst_tac [("z","z2")] eq_Abs_prat 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   219
by (res_inst_tac [("z","z3")] eq_Abs_prat 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   220
by (asm_simp_tac (simpset() addsimps [prat_mult, pnat_mult_assoc]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   221
qed "prat_mult_assoc";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   222
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   223
(*For AC rewriting*)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   224
Goal "(x::prat)*(y*z)=y*(x*z)";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   225
by (rtac (prat_mult_commute RS trans) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   226
by (rtac (prat_mult_assoc RS trans) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   227
by (rtac (prat_mult_commute RS arg_cong) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   228
qed "prat_mult_left_commute";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   229
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   230
(*Positive Rational multiplication is an AC operator*)
9108
9fff97d29837 bind_thm(s);
wenzelm
parents: 9043
diff changeset
   231
bind_thms ("prat_mult_ac", [prat_mult_assoc,
9fff97d29837 bind_thm(s);
wenzelm
parents: 9043
diff changeset
   232
                    prat_mult_commute,prat_mult_left_commute]);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   233
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
   234
Goalw [prat_of_pnat_def] 
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11655
diff changeset
   235
      "(prat_of_pnat (Abs_pnat (Suc 0))) * z = z";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   236
by (res_inst_tac [("z","z")] eq_Abs_prat 1);
7376
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   237
by (asm_full_simp_tac (simpset() addsimps [prat_mult] @ pnat_mult_ac) 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   238
qed "prat_mult_1";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   239
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
   240
Goalw [prat_of_pnat_def] 
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11655
diff changeset
   241
      "z * (prat_of_pnat (Abs_pnat (Suc 0))) = z";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   242
by (res_inst_tac [("z","z")] eq_Abs_prat 1);
7376
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   243
by (asm_full_simp_tac (simpset() addsimps [prat_mult] @ pnat_mult_ac) 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   244
qed "prat_mult_1_right";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   245
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
   246
Goalw [prat_of_pnat_def] 
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
   247
      "prat_of_pnat ((z1::pnat) + z2) = \
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
   248
\      prat_of_pnat z1 + prat_of_pnat z2";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   249
by (asm_simp_tac (simpset() addsimps [prat_add,
7376
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   250
				      pnat_add_mult_distrib,pnat_mult_1]) 1);
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
   251
qed "prat_of_pnat_add";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   252
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
   253
Goalw [prat_of_pnat_def] 
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
   254
      "prat_of_pnat ((z1::pnat) * z2) = \
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
   255
\      prat_of_pnat z1 * prat_of_pnat z2";
7376
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   256
by (asm_simp_tac (simpset() addsimps [prat_mult, pnat_mult_1]) 1);
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
   257
qed "prat_of_pnat_mult";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   258
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   259
(*** prat_mult and qinv ***)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   260
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
   261
Goalw [prat_def,prat_of_pnat_def] 
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11655
diff changeset
   262
      "qinv (q) * q = prat_of_pnat (Abs_pnat (Suc 0))";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   263
by (res_inst_tac [("z","q")] eq_Abs_prat 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   264
by (asm_full_simp_tac (simpset() addsimps [qinv,
7376
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   265
        prat_mult,pnat_mult_1,pnat_mult_1_left, pnat_mult_commute]) 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   266
qed "prat_mult_qinv";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   267
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11655
diff changeset
   268
Goal "q * qinv (q) = prat_of_pnat (Abs_pnat (Suc 0))";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   269
by (rtac (prat_mult_commute RS subst) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   270
by (simp_tac (simpset() addsimps [prat_mult_qinv]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   271
qed "prat_mult_qinv_right";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   272
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11655
diff changeset
   273
Goal "EX y. (x::prat) * y = prat_of_pnat (Abs_pnat (Suc 0))";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   274
by (fast_tac (claset() addIs [prat_mult_qinv_right]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   275
qed "prat_qinv_ex";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   276
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11655
diff changeset
   277
Goal "EX! y. (x::prat) * y = prat_of_pnat (Abs_pnat (Suc 0))";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   278
by (auto_tac (claset() addIs [prat_mult_qinv_right],simpset()));
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   279
by (dres_inst_tac [("f","%x. ya*x")] arg_cong 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   280
by (asm_full_simp_tac (simpset() addsimps [prat_mult_assoc RS sym]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   281
by (asm_full_simp_tac (simpset() addsimps [prat_mult_commute,
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   282
    prat_mult_1,prat_mult_1_right]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   283
qed "prat_qinv_ex1";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   284
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11655
diff changeset
   285
Goal "EX! y. y * (x::prat) = prat_of_pnat (Abs_pnat (Suc 0))";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   286
by (auto_tac (claset() addIs [prat_mult_qinv],simpset()));
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   287
by (dres_inst_tac [("f","%x. x*ya")] arg_cong 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   288
by (asm_full_simp_tac (simpset() addsimps [prat_mult_assoc]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   289
by (asm_full_simp_tac (simpset() addsimps [prat_mult_commute,
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   290
    prat_mult_1,prat_mult_1_right]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   291
qed "prat_qinv_left_ex1";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   292
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11655
diff changeset
   293
Goal "x * y = prat_of_pnat (Abs_pnat (Suc 0)) ==> x = qinv y";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   294
by (cut_inst_tac [("q","y")] prat_mult_qinv 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   295
by (res_inst_tac [("x1","y")] (prat_qinv_left_ex1 RS ex1E) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   296
by (Blast_tac 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   297
qed "prat_mult_inv_qinv";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   298
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 7825
diff changeset
   299
Goal "EX y. x = qinv y";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   300
by (cut_inst_tac [("x","x")] prat_qinv_ex 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   301
by (etac exE 1 THEN dtac prat_mult_inv_qinv 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   302
by (Fast_tac 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   303
qed "prat_as_inverse_ex";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   304
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   305
Goal "qinv(x*y) = qinv(x)*qinv(y)";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   306
by (res_inst_tac [("z","x")] eq_Abs_prat 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   307
by (res_inst_tac [("z","y")] eq_Abs_prat 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   308
by (auto_tac (claset(),simpset() addsimps [qinv,prat_mult]));
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   309
qed "qinv_mult_eq";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   310
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   311
(** Lemmas **)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   312
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   313
Goal "((z1::prat) + z2) * w = (z1 * w) + (z2 * w)";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   314
by (res_inst_tac [("z","z1")] eq_Abs_prat 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   315
by (res_inst_tac [("z","z2")] eq_Abs_prat 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   316
by (res_inst_tac [("z","w")] eq_Abs_prat 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   317
by (asm_simp_tac 
5535
678999604ee9 deleted needless parentheses
paulson
parents: 5459
diff changeset
   318
    (simpset() addsimps [pnat_add_mult_distrib2, prat_add, prat_mult] @ 
678999604ee9 deleted needless parentheses
paulson
parents: 5459
diff changeset
   319
                        pnat_add_ac @ pnat_mult_ac) 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   320
qed "prat_add_mult_distrib";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   321
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   322
val prat_mult_commute'= read_instantiate [("z","w")] prat_mult_commute;
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   323
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   324
Goal "(w::prat) * (z1 + z2) = (w * z1) + (w * z2)";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   325
by (simp_tac (simpset() addsimps [prat_mult_commute',prat_add_mult_distrib]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   326
qed "prat_add_mult_distrib2";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   327
7376
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   328
Addsimps [prat_mult_1, prat_mult_1_right, 
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   329
	  prat_mult_qinv, prat_mult_qinv_right];
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   330
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   331
      (*** theorems for ordering ***)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   332
(* prove introduction and elimination rules for prat_less *)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   333
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   334
Goalw [prat_less_def]
11655
923e4d0d36d5 tuned parentheses in relational expressions;
wenzelm
parents: 11464
diff changeset
   335
    "(Q1 < (Q2::prat)) = (EX Q3. Q1 + Q3 = Q2)";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   336
by (Fast_tac 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   337
qed "prat_less_iff";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   338
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   339
Goalw [prat_less_def]
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   340
      "!!(Q1::prat). Q1 + Q3 = Q2 ==> Q1 < Q2";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   341
by (Fast_tac  1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   342
qed "prat_lessI";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   343
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   344
(* ordering on positive fractions in terms of existence of sum *)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   345
Goalw [prat_less_def]
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   346
      "Q1 < (Q2::prat) --> (EX Q3. Q1 + Q3 = Q2)";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   347
by (Fast_tac 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   348
qed "prat_lessE_lemma";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   349
5148
74919e8f221c More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5143
diff changeset
   350
Goal "!!P. [| Q1 < (Q2::prat); \
74919e8f221c More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5143
diff changeset
   351
\             !! (Q3::prat). Q1 + Q3 = Q2 ==> P |] \
74919e8f221c More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5143
diff changeset
   352
\          ==> P";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   353
by (dtac (prat_lessE_lemma RS mp) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   354
by Auto_tac;
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   355
qed "prat_lessE";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   356
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   357
(* qless is a strong order i.e nonreflexive and transitive *)
5148
74919e8f221c More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5143
diff changeset
   358
Goal "!!(q1::prat). [| q1 < q2; q2 < q3 |] ==> q1 < q3";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   359
by (REPEAT(dtac (prat_lessE_lemma RS mp) 1));
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   360
by (REPEAT(etac exE 1));
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   361
by (hyp_subst_tac 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   362
by (res_inst_tac [("Q3.0","Q3 + Q3a")] prat_lessI 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   363
by (auto_tac (claset(),simpset() addsimps [prat_add_assoc]));
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   364
qed "prat_less_trans";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   365
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   366
Goal "~q < (q::prat)";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   367
by (EVERY1[rtac notI, dtac (prat_lessE_lemma RS mp)]);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   368
by (res_inst_tac [("z","q")] eq_Abs_prat 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   369
by (res_inst_tac [("z","Q3")] eq_Abs_prat 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   370
by (etac exE 1 THEN res_inst_tac [("z","Q3a")] eq_Abs_prat 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   371
by (REPEAT(hyp_subst_tac 1));
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   372
by (asm_full_simp_tac (simpset() addsimps [prat_add,
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   373
    pnat_no_add_ident,pnat_add_mult_distrib2] @ pnat_mult_ac) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   374
qed "prat_less_not_refl";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   375
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   376
(*** y < y ==> P ***)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   377
bind_thm("prat_less_irrefl",prat_less_not_refl RS notE);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   378
5459
1dbaf888f4e7 well-formed asym rules; misc. tidying
paulson
parents: 5148
diff changeset
   379
Goal "!! (q1::prat). q1 < q2 ==> ~ q2 < q1";
1dbaf888f4e7 well-formed asym rules; misc. tidying
paulson
parents: 5148
diff changeset
   380
by (rtac notI 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   381
by (dtac prat_less_trans 1 THEN assume_tac 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   382
by (asm_full_simp_tac (simpset() addsimps [prat_less_not_refl]) 1);
5459
1dbaf888f4e7 well-formed asym rules; misc. tidying
paulson
parents: 5148
diff changeset
   383
qed "prat_less_not_sym";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   384
5459
1dbaf888f4e7 well-formed asym rules; misc. tidying
paulson
parents: 5148
diff changeset
   385
(* [| x < y;  ~P ==> y < x |] ==> P *)
10232
529c65b5dcde restoration of "equalityI"; renaming of contrapos rules
paulson
parents: 9969
diff changeset
   386
bind_thm ("prat_less_asym", prat_less_not_sym RS contrapos_np);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   387
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   388
(* half of positive fraction exists- Gleason p. 120- Proposition 9-2.6(i)*)
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 7825
diff changeset
   389
Goal "!(q::prat). EX x. x + x = q";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   390
by (rtac allI 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   391
by (res_inst_tac [("z","q")] eq_Abs_prat 1);
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
   392
by (res_inst_tac [("x","Abs_prat (ratrel `` {(x, y+y)})")] exI 1);
7376
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   393
by (auto_tac (claset(),
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   394
	      simpset() addsimps 
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   395
              [prat_add,pnat_mult_assoc RS sym,pnat_add_mult_distrib,
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   396
               pnat_add_mult_distrib2]));
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   397
qed "lemma_prat_dense";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   398
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 7825
diff changeset
   399
Goal "EX (x::prat). x + x = q";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   400
by (res_inst_tac [("z","q")] eq_Abs_prat 1);
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
   401
by (res_inst_tac [("x","Abs_prat (ratrel `` {(x, y+y)})")] exI 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   402
by (auto_tac (claset(),simpset() addsimps 
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   403
              [prat_add,pnat_mult_assoc RS sym,pnat_add_mult_distrib,
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   404
               pnat_add_mult_distrib2]));
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   405
qed "prat_lemma_dense";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   406
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   407
(* there exists a number between any two positive fractions *)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   408
(* Gleason p. 120- Proposition 9-2.6(iv) *)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   409
Goalw [prat_less_def] 
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 7825
diff changeset
   410
      "!! (q1::prat). q1 < q2 ==> EX x. q1 < x & x < q2";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   411
by (auto_tac (claset() addIs [lemma_prat_dense],simpset()));
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   412
by (res_inst_tac [("x","T")] (lemma_prat_dense RS allE) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   413
by (etac exE 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   414
by (res_inst_tac [("x","q1 + x")] exI 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   415
by (auto_tac (claset() addIs [prat_lemma_dense],
7376
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   416
	      simpset() addsimps [prat_add_assoc]));
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   417
qed "prat_dense";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   418
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   419
(* ordering of addition for positive fractions *)
7376
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   420
Goalw [prat_less_def] "!!(q1::prat). q1 < q2 ==> q1 + x < q2 + x";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   421
by (Step_tac 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   422
by (res_inst_tac [("x","T")] exI 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   423
by (auto_tac (claset(),simpset() addsimps prat_add_ac));
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   424
qed "prat_add_less2_mono1";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   425
7219
4e3f386c2e37 inserted Id: lines
paulson
parents: 7077
diff changeset
   426
Goal "!!(q1::prat). q1 < q2 ==> x + q1 < x + q2";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   427
by (auto_tac (claset() addIs [prat_add_less2_mono1],
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   428
    simpset() addsimps [prat_add_commute]));
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   429
qed "prat_add_less2_mono2";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   430
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   431
(* ordering of multiplication for positive fractions *)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   432
Goalw [prat_less_def] 
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   433
      "!!(q1::prat). q1 < q2 ==> q1 * x < q2 * x";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   434
by (Step_tac 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   435
by (res_inst_tac [("x","T*x")] exI 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   436
by (auto_tac (claset(),simpset() addsimps [prat_add_mult_distrib]));
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   437
qed "prat_mult_less2_mono1";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   438
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   439
Goal "!!(q1::prat). q1 < q2  ==> x * q1 < x * q2";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   440
by (auto_tac (claset() addDs [prat_mult_less2_mono1],
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   441
    simpset() addsimps [prat_mult_commute]));
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   442
qed "prat_mult_left_less2_mono1";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   443
9426
d29faa6cc391 avoid referencing thy value;
wenzelm
parents: 9391
diff changeset
   444
Goal "!!(a1::prat). a1 < a2 ==> a1 * b + a2 * c < a2 * (b + c)";
d29faa6cc391 avoid referencing thy value;
wenzelm
parents: 9391
diff changeset
   445
by (auto_tac (claset() addSIs [prat_add_less2_mono1,prat_mult_less2_mono1],
d29faa6cc391 avoid referencing thy value;
wenzelm
parents: 9391
diff changeset
   446
              simpset() addsimps [prat_add_mult_distrib2]));
d29faa6cc391 avoid referencing thy value;
wenzelm
parents: 9391
diff changeset
   447
qed "lemma_prat_add_mult_mono";
d29faa6cc391 avoid referencing thy value;
wenzelm
parents: 9391
diff changeset
   448
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   449
(* there is no smallest positive fraction *)
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 7825
diff changeset
   450
Goalw [prat_less_def] "EX (x::prat). x < y";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   451
by (cut_facts_tac [lemma_prat_dense] 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   452
by (Fast_tac 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   453
qed "qless_Ex";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   454
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   455
(* lemma for proving $< is linear *)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   456
Goalw [prat_def,prat_less_def] 
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
   457
      "ratrel `` {(x, y * ya)} : {p::(pnat*pnat).True}//ratrel";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   458
by (asm_full_simp_tac (simpset() addsimps [ratrel_def,quotient_def]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   459
by (Blast_tac 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   460
qed "lemma_prat_less_linear";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   461
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   462
(* linearity of < -- Gleason p. 120 - Proposition 9-2.6 *)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   463
(*** FIXME Proof long ***)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   464
Goalw [prat_less_def] 
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   465
      "(q1::prat) < q2 | q1 = q2 | q2 < q1";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   466
by (res_inst_tac [("z","q1")] eq_Abs_prat 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   467
by (res_inst_tac [("z","q2")] eq_Abs_prat 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   468
by (Step_tac 1 THEN REPEAT(dtac (not_ex RS iffD1) 1) 
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   469
               THEN Auto_tac);
7376
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   470
by (cut_inst_tac  [("z1.0","x*ya"), ("z2.0","xa*y")] pnat_linear_Ex_eq 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   471
by (EVERY1[etac disjE,etac exE]);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   472
by (eres_inst_tac 
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
   473
    [("x","Abs_prat(ratrel``{(xb,ya*y)})")] allE 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   474
by (asm_full_simp_tac 
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   475
    (simpset() addsimps [prat_add, pnat_mult_assoc 
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   476
     RS sym,pnat_add_mult_distrib RS sym]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   477
by (EVERY1[asm_full_simp_tac (simpset() addsimps pnat_mult_ac),
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   478
    etac disjE, assume_tac, etac exE]);
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
   479
by (thin_tac "!T. Abs_prat (ratrel `` {(x, y)}) + T ~= \
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
   480
\     Abs_prat (ratrel `` {(xa, ya)})" 1);
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
   481
by (eres_inst_tac [("x","Abs_prat(ratrel``{(xb,y*ya)})")] allE 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   482
by (asm_full_simp_tac (simpset() addsimps [prat_add,
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   483
      pnat_mult_assoc RS sym,pnat_add_mult_distrib RS sym]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   484
by (asm_full_simp_tac (simpset() addsimps pnat_mult_ac) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   485
qed "prat_linear";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   486
7219
4e3f386c2e37 inserted Id: lines
paulson
parents: 7077
diff changeset
   487
Goal "!!(q1::prat). [| q1 < q2 ==> P;  q1 = q2 ==> P; \
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   488
\          q2 < q1 ==> P |] ==> P";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   489
by (cut_inst_tac [("q1.0","q1"),("q2.0","q2")] prat_linear 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   490
by Auto_tac;
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   491
qed "prat_linear_less2";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   492
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   493
(* Gleason p. 120 -- 9-2.6 (iv) *)
7376
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   494
Goal "[| q1 < q2; qinv(q1) = qinv(q2) |] ==> P";
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   495
by (cut_inst_tac [("x","qinv (q2)"),("q1.0","q1"), ("q2.0","q2")] 
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   496
    prat_mult_less2_mono1 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   497
by (assume_tac 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   498
by (Asm_full_simp_tac 1 THEN dtac sym 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   499
by (auto_tac (claset(),simpset() addsimps [prat_less_not_refl]));
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   500
qed "lemma1_qinv_prat_less";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   501
7376
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   502
Goal "[| q1 < q2; qinv(q1) < qinv(q2) |] ==> P";
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   503
by (cut_inst_tac [("x","qinv (q2)"),("q1.0","q1"), ("q2.0","q2")] 
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   504
    prat_mult_less2_mono1 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   505
by (assume_tac 1);
7376
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   506
by (cut_inst_tac [("x","q1"),("q1.0","qinv (q1)"), ("q2.0","qinv (q2)")] 
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   507
    prat_mult_left_less2_mono1 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   508
by Auto_tac;
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11655
diff changeset
   509
by (dres_inst_tac [("q2.0","prat_of_pnat (Abs_pnat (Suc 0))")] prat_less_trans 1);
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
   510
by (auto_tac (claset(),simpset() addsimps 
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
   511
    [prat_less_not_refl]));
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   512
qed "lemma2_qinv_prat_less";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   513
7376
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   514
Goal "q1 < q2  ==> qinv (q2) < qinv (q1)";
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   515
by (res_inst_tac [("q2.0","qinv q1"), ("q1.0","qinv q2")] prat_linear_less2 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   516
by (auto_tac (claset() addEs [lemma1_qinv_prat_less,
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   517
                 lemma2_qinv_prat_less],simpset()));
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   518
qed "qinv_prat_less";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   519
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11655
diff changeset
   520
Goal "q1 < prat_of_pnat (Abs_pnat (Suc 0)) \
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11655
diff changeset
   521
\     ==> prat_of_pnat (Abs_pnat (Suc 0)) < qinv(q1)";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   522
by (dtac qinv_prat_less 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   523
by (full_simp_tac (simpset() addsimps [qinv_1]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   524
qed "prat_qinv_gt_1";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   525
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
   526
Goalw [pnat_one_def] 
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11701
diff changeset
   527
     "q1 < prat_of_pnat 1 ==> prat_of_pnat 1 < qinv(q1)";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   528
by (etac prat_qinv_gt_1 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   529
qed "prat_qinv_is_gt_1";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   530
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
   531
Goalw [prat_less_def] 
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11655
diff changeset
   532
      "prat_of_pnat (Abs_pnat (Suc 0)) < prat_of_pnat (Abs_pnat (Suc 0)) \
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11655
diff changeset
   533
\                   + prat_of_pnat (Abs_pnat (Suc 0))";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   534
by (Fast_tac 1); 
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   535
qed "prat_less_1_2";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   536
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11655
diff changeset
   537
Goal "qinv(prat_of_pnat (Abs_pnat (Suc 0)) + \
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11655
diff changeset
   538
\     prat_of_pnat (Abs_pnat (Suc 0))) < prat_of_pnat (Abs_pnat (Suc 0))";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   539
by (cut_facts_tac [prat_less_1_2 RS qinv_prat_less] 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   540
by (asm_full_simp_tac (simpset() addsimps [qinv_1]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   541
qed "prat_less_qinv_2_1";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   542
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11655
diff changeset
   543
Goal "!!(x::prat). x < y ==> x*qinv(y) < prat_of_pnat (Abs_pnat (Suc 0))";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   544
by (dres_inst_tac [("x","qinv(y)")] prat_mult_less2_mono1 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   545
by (Asm_full_simp_tac 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   546
qed "prat_mult_qinv_less_1";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   547
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   548
Goal "(x::prat) < x + x";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   549
by (cut_inst_tac [("x","x")] 
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   550
    (prat_less_1_2 RS prat_mult_left_less2_mono1) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   551
by (asm_full_simp_tac (simpset() addsimps 
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   552
    [prat_add_mult_distrib2]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   553
qed "prat_self_less_add_self";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   554
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   555
Goalw [prat_less_def] "(x::prat) < y + x";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   556
by (res_inst_tac [("x","y")] exI 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   557
by (simp_tac (simpset() addsimps [prat_add_commute]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   558
qed "prat_self_less_add_right";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   559
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   560
Goal "(x::prat) < x + y";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   561
by (rtac (prat_add_commute RS subst) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   562
by (simp_tac (simpset() addsimps [prat_self_less_add_right]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   563
qed "prat_self_less_add_left";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   564
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11701
diff changeset
   565
Goalw [prat_less_def] "prat_of_pnat 1 < y ==> (x::prat) < x * y";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   566
by (auto_tac (claset(),simpset() addsimps [pnat_one_def,
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   567
    prat_add_mult_distrib2]));
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   568
qed "prat_self_less_mult_right";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   569
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   570
(*** Properties of <= ***)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   571
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5078
diff changeset
   572
Goalw [prat_le_def] "~(w < z) ==> z <= (w::prat)";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   573
by (assume_tac 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   574
qed "prat_leI";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   575
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5078
diff changeset
   576
Goalw [prat_le_def] "z<=w ==> ~(w<(z::prat))";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   577
by (assume_tac 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   578
qed "prat_leD";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   579
9108
9fff97d29837 bind_thm(s);
wenzelm
parents: 9043
diff changeset
   580
bind_thm ("prat_leE", make_elim prat_leD);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   581
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5078
diff changeset
   582
Goal "(~(w < z)) = (z <= (w::prat))";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   583
by (fast_tac (claset() addSIs [prat_leI,prat_leD]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   584
qed "prat_less_le_iff";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   585
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5078
diff changeset
   586
Goalw [prat_le_def] "~ z <= w ==> w<(z::prat)";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   587
by (Fast_tac 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   588
qed "not_prat_leE";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   589
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5078
diff changeset
   590
Goalw [prat_le_def] "z < w ==> z <= (w::prat)";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   591
by (fast_tac (claset() addEs [prat_less_asym]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   592
qed "prat_less_imp_le";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   593
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   594
Goalw [prat_le_def] "!!(x::prat). x <= y ==> x < y | x = y";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   595
by (cut_facts_tac [prat_linear] 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   596
by (fast_tac (claset() addEs [prat_less_irrefl,prat_less_asym]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   597
qed "prat_le_imp_less_or_eq";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   598
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5078
diff changeset
   599
Goalw [prat_le_def] "z<w | z=w ==> z <=(w::prat)";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   600
by (cut_facts_tac [prat_linear] 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   601
by (fast_tac (claset() addEs [prat_less_irrefl,prat_less_asym]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   602
qed "prat_less_or_eq_imp_le";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   603
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   604
Goal "(x <= (y::prat)) = (x < y | x=y)";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   605
by (REPEAT(ares_tac [iffI, prat_less_or_eq_imp_le, prat_le_imp_less_or_eq] 1));
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   606
qed "prat_le_eq_less_or_eq";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   607
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   608
Goal "w <= (w::prat)";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   609
by (simp_tac (simpset() addsimps [prat_le_eq_less_or_eq]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   610
qed "prat_le_refl";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   611
9969
4753185f1dd2 renamed (most of...) the select rules
paulson
parents: 9426
diff changeset
   612
Goal "[| i <= j; j < k |] ==> i < (k::prat)";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   613
by (dtac prat_le_imp_less_or_eq 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   614
by (fast_tac (claset() addIs [prat_less_trans]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   615
qed "prat_le_less_trans";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   616
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5078
diff changeset
   617
Goal "[| i <= j; j <= k |] ==> i <= (k::prat)";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   618
by (EVERY1 [dtac prat_le_imp_less_or_eq, dtac prat_le_imp_less_or_eq,
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   619
            rtac prat_less_or_eq_imp_le, fast_tac (claset() addIs [prat_less_trans])]);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   620
qed "prat_le_trans";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   621
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5078
diff changeset
   622
Goal "[| ~ y < x; y ~= x |] ==> x < (y::prat)";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   623
by (rtac not_prat_leE 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   624
by (fast_tac (claset() addDs [prat_le_imp_less_or_eq]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   625
qed "not_less_not_eq_prat_less";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   626
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   627
Goalw [prat_less_def] 
5148
74919e8f221c More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5143
diff changeset
   628
      "[| x1 < y1; x2 < y2 |] ==> x1 + x2 < y1 + (y2::prat)";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   629
by (REPEAT(etac exE 1));
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   630
by (res_inst_tac [("x","T+Ta")] exI 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   631
by (auto_tac (claset(),simpset() addsimps prat_add_ac));
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   632
qed "prat_add_less_mono";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   633
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   634
Goalw [prat_less_def] 
5148
74919e8f221c More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5143
diff changeset
   635
      "[| x1 < y1; x2 < y2 |] ==> x1 * x2 < y1 * (y2::prat)";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   636
by (REPEAT(etac exE 1));
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   637
by (res_inst_tac [("x","T*Ta+T*x2+x1*Ta")] exI 1);
7376
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   638
by (auto_tac (claset(),
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   639
	      simpset() addsimps prat_add_ac @ 
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   640
	                      [prat_add_mult_distrib,prat_add_mult_distrib2]));
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   641
qed "prat_mult_less_mono";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   642
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   643
(* more prat_le *)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   644
Goal "!!(q1::prat). q1 <= q2  ==> x * q1 <= x * q2";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   645
by (dtac prat_le_imp_less_or_eq 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   646
by (Step_tac 1);
7376
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   647
by (auto_tac (claset() addSIs [prat_le_refl, prat_less_imp_le,
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   648
			       prat_mult_left_less2_mono1],
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   649
	      simpset()));
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   650
qed "prat_mult_left_le2_mono1";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   651
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   652
Goal "!!(q1::prat). q1 <= q2  ==> q1 * x <= q2 * x";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   653
by (auto_tac (claset() addDs [prat_mult_left_le2_mono1],
7376
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   654
	      simpset() addsimps [prat_mult_commute]));
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   655
qed "prat_mult_le2_mono1";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   656
7376
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   657
Goal "q1 <= q2  ==> qinv (q2) <= qinv (q1)";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   658
by (dtac prat_le_imp_less_or_eq 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   659
by (Step_tac 1);
7376
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   660
by (auto_tac (claset() addSIs [prat_le_refl, prat_less_imp_le,qinv_prat_less],
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   661
	      simpset()));
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   662
qed "qinv_prat_le";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   663
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   664
Goal "!!(q1::prat). q1 <= q2  ==> x + q1 <= x + q2";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   665
by (dtac prat_le_imp_less_or_eq 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   666
by (Step_tac 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   667
by (auto_tac (claset() addSIs [prat_le_refl,
7376
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   668
			       prat_less_imp_le,prat_add_less2_mono1],
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   669
	      simpset() addsimps [prat_add_commute]));
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   670
qed "prat_add_left_le2_mono1";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   671
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   672
Goal "!!(q1::prat). q1 <= q2  ==> q1 + x <= q2 + x";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   673
by (auto_tac (claset() addDs [prat_add_left_le2_mono1],
7376
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   674
	      simpset() addsimps [prat_add_commute]));
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   675
qed "prat_add_le2_mono1";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   676
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   677
Goal "!!k l::prat. [|i<=j;  k<=l |] ==> i + k <= j + l";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   678
by (etac (prat_add_le2_mono1 RS prat_le_trans) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   679
by (simp_tac (simpset() addsimps [prat_add_commute]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   680
(*j moves to the end because it is free while k, l are bound*)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   681
by (etac prat_add_le2_mono1 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   682
qed "prat_add_le_mono";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   683
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   684
Goal "!!(x::prat). x + y < z + y ==> x < z";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   685
by (rtac ccontr 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   686
by (etac (prat_leI RS prat_le_imp_less_or_eq RS disjE) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   687
by (dres_inst_tac [("x","y"),("q1.0","z")] prat_add_less2_mono1 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   688
by (auto_tac (claset() addIs [prat_less_asym],
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   689
    simpset() addsimps [prat_less_not_refl]));
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   690
qed "prat_add_right_less_cancel";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   691
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   692
Goal "!!(x::prat). y + x < y + z ==> x < z";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   693
by (res_inst_tac [("y","y")] prat_add_right_less_cancel 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   694
by (asm_full_simp_tac (simpset() addsimps [prat_add_commute]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   695
qed "prat_add_left_less_cancel";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   696
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   697
(*** lemmas required for lemma_gleason9_34 in PReal : w*y > y/z ***)
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
   698
Goalw [prat_of_pnat_def] 
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
   699
      "Abs_prat(ratrel``{(x,y)}) = (prat_of_pnat x)*qinv(prat_of_pnat y)";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   700
by (auto_tac (claset(),simpset() addsimps [prat_mult,qinv,pnat_mult_1_left,
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   701
    pnat_mult_1]));
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   702
qed "Abs_prat_mult_qinv";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   703
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11655
diff changeset
   704
Goal "Abs_prat(ratrel``{(x,y)}) <= Abs_prat(ratrel``{(x,Abs_pnat (Suc 0))})";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   705
by (simp_tac (simpset() addsimps [Abs_prat_mult_qinv]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   706
by (rtac prat_mult_left_le2_mono1 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   707
by (rtac qinv_prat_le 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   708
by (pnat_ind_tac "y" 1);
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11655
diff changeset
   709
by (dres_inst_tac [("x","prat_of_pnat (Abs_pnat (Suc 0))")] prat_add_le2_mono1 2);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   710
by (cut_facts_tac [prat_less_1_2 RS prat_less_imp_le] 2);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   711
by (auto_tac (claset() addIs [prat_le_trans],
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   712
    simpset() addsimps [prat_le_refl,
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
   713
    pSuc_is_plus_one,pnat_one_def,prat_of_pnat_add]));
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   714
qed "lemma_Abs_prat_le1";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   715
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11655
diff changeset
   716
Goal "Abs_prat(ratrel``{(x,Abs_pnat (Suc 0))}) <= Abs_prat(ratrel``{(x*y,Abs_pnat (Suc 0))})";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   717
by (simp_tac (simpset() addsimps [Abs_prat_mult_qinv]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   718
by (rtac prat_mult_le2_mono1 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   719
by (pnat_ind_tac "y" 1);
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
   720
by (dres_inst_tac [("x","prat_of_pnat x")] prat_add_le2_mono1 2);
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
   721
by (cut_inst_tac [("z","prat_of_pnat x")] (prat_self_less_add_self 
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   722
    RS prat_less_imp_le) 2);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   723
by (auto_tac (claset() addIs [prat_le_trans],
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   724
    simpset() addsimps [prat_le_refl,
7376
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   725
			pSuc_is_plus_one,pnat_one_def,prat_add_mult_distrib2,
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   726
			prat_of_pnat_add,prat_of_pnat_mult]));
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   727
qed "lemma_Abs_prat_le2";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   728
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11655
diff changeset
   729
Goal "Abs_prat(ratrel``{(x,z)}) <= Abs_prat(ratrel``{(x*y,Abs_pnat (Suc 0))})";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
   730
by (fast_tac (claset() addIs [prat_le_trans,
7376
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   731
			      lemma_Abs_prat_le1,lemma_Abs_prat_le2]) 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   732
qed "lemma_Abs_prat_le3";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   733
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11655
diff changeset
   734
Goal "Abs_prat(ratrel``{(x*y,Abs_pnat (Suc 0))}) * Abs_prat(ratrel``{(w,x)}) = \
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11655
diff changeset
   735
\         Abs_prat(ratrel``{(w*y,Abs_pnat (Suc 0))})";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   736
by (full_simp_tac (simpset() addsimps [prat_mult,
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   737
    pnat_mult_1,pnat_mult_1_left] @ pnat_mult_ac) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   738
qed "pre_lemma_gleason9_34";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   739
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11655
diff changeset
   740
Goal "Abs_prat(ratrel``{(y*x,Abs_pnat (Suc 0)*y)}) = \
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11655
diff changeset
   741
\         Abs_prat(ratrel``{(x,Abs_pnat (Suc 0))})";
7376
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   742
by (auto_tac (claset(),
46f92a120af9 tidied, allowing pattern-matching in defs of prat_add and prat_mult
paulson
parents: 7219
diff changeset
   743
	      simpset() addsimps [pnat_mult_1,pnat_mult_1_left] @ pnat_mult_ac));
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   744
qed "pre_lemma_gleason9_34b";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   745
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
   746
Goal "(prat_of_pnat n < prat_of_pnat m) = (n < m)";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   747
by (auto_tac (claset(),simpset() addsimps [prat_less_def,
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
   748
    pnat_less_iff,prat_of_pnat_add]));
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   749
by (res_inst_tac [("z","T")] eq_Abs_prat 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   750
by (auto_tac (claset() addDs [pnat_eq_lessI],
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   751
    simpset() addsimps [prat_add,pnat_mult_1,
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
   752
    pnat_mult_1_left,prat_of_pnat_def,pnat_less_iff RS sym]));
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
   753
qed "prat_of_pnat_less_iff";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   754
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
   755
Addsimps [prat_of_pnat_less_iff];
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   756
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
   757
(*------------------------------------------------------------------*)
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   758
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   759
(*** prove witness that will be required to prove non-emptiness ***)
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
   760
(*** of preal type as defined using Dedekind Sections in PReal  ***)
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   761
(*** Show that exists positive real `one' ***)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   762
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11655
diff changeset
   763
Goal "EX q. q: {x::prat. x < prat_of_pnat (Abs_pnat (Suc 0))}";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   764
by (fast_tac (claset() addIs [prat_less_qinv_2_1]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   765
qed "lemma_prat_less_1_memEx";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   766
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11655
diff changeset
   767
Goal "{x::prat. x < prat_of_pnat (Abs_pnat (Suc 0))} ~= {}";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   768
by (rtac notI 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   769
by (cut_facts_tac [lemma_prat_less_1_memEx] 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   770
by (Asm_full_simp_tac 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   771
qed "lemma_prat_less_1_set_non_empty";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   772
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11655
diff changeset
   773
Goalw [psubset_def] "{} < {x::prat. x < prat_of_pnat (Abs_pnat (Suc 0))}";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   774
by (asm_full_simp_tac (simpset() addsimps 
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   775
         [lemma_prat_less_1_set_non_empty RS not_sym]) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   776
qed "empty_set_psubset_lemma_prat_less_1_set";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   777
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
   778
(*** exists rational not in set --- prat_of_pnat (Abs_pnat 1) itself ***)
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11655
diff changeset
   779
Goal "EX q. q ~: {x::prat. x < prat_of_pnat (Abs_pnat (Suc 0))}";
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11655
diff changeset
   780
by (res_inst_tac [("x","prat_of_pnat (Abs_pnat (Suc 0))")] exI 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   781
by (auto_tac (claset(),simpset() addsimps [prat_less_not_refl]));
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   782
qed "lemma_prat_less_1_not_memEx";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   783
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11655
diff changeset
   784
Goal "{x::prat. x < prat_of_pnat (Abs_pnat (Suc 0))} ~= UNIV";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   785
by (rtac notI 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   786
by (cut_facts_tac [lemma_prat_less_1_not_memEx] 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   787
by (Asm_full_simp_tac 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   788
qed "lemma_prat_less_1_set_not_rat_set";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   789
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   790
Goalw [psubset_def,subset_def] 
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11655
diff changeset
   791
      "{x::prat. x < prat_of_pnat (Abs_pnat (Suc 0))} < UNIV";
7825
1be9b63e7d93 replaced {x. True} by UNIV to work with the new simprule, Collect_const
paulson
parents: 7376
diff changeset
   792
by (asm_full_simp_tac
1be9b63e7d93 replaced {x. True} by UNIV to work with the new simprule, Collect_const
paulson
parents: 7376
diff changeset
   793
    (simpset() addsimps [lemma_prat_less_1_set_not_rat_set,
1be9b63e7d93 replaced {x. True} by UNIV to work with the new simprule, Collect_const
paulson
parents: 7376
diff changeset
   794
			 lemma_prat_less_1_not_memEx]) 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   795
qed "lemma_prat_less_1_set_psubset_rat_set";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   796
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   797
(*** prove non_emptiness of type ***)
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11655
diff changeset
   798
Goal "{x::prat. x < prat_of_pnat (Abs_pnat (Suc 0))} : {A. {} < A & \
7825
1be9b63e7d93 replaced {x. True} by UNIV to work with the new simprule, Collect_const
paulson
parents: 7376
diff changeset
   799
\               A < UNIV & \
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5535
diff changeset
   800
\               (!y: A. ((!z. z < y --> z: A) & \
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 7825
diff changeset
   801
\               (EX u: A. y < u)))}";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   802
by (auto_tac (claset() addDs [prat_less_trans],
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   803
    simpset() addsimps [empty_set_psubset_lemma_prat_less_1_set,
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   804
                       lemma_prat_less_1_set_psubset_rat_set]));
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   805
by (dtac prat_dense 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   806
by (Fast_tac 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   807
qed "preal_1";