src/HOL/NatDef.ML
author paulson
Fri, 18 Feb 2000 15:33:09 +0100
changeset 8253 975eb12aa040
parent 7064 b053e0ab9f60
child 8555 17325ee838ab
permissions -rw-r--r--
many new theorems about inj, surj etc.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/NatDef.ML
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
     2
    ID:         $Id$
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
     3
    Author:     Tobias Nipkow, Cambridge University Computer Laboratory
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
     4
    Copyright   1991  University of Cambridge
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
     5
*)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
     6
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
     7
Goal "mono(%X. {Zero_Rep} Un (Suc_Rep``X))";
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
     8
by (REPEAT (ares_tac [monoI, subset_refl, image_mono, Un_mono] 1));
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
     9
qed "Nat_fun_mono";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    10
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    11
val Nat_unfold = Nat_fun_mono RS (Nat_def RS def_lfp_Tarski);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    12
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    13
(* Zero is a natural number -- this also justifies the type definition*)
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
    14
Goal "Zero_Rep: Nat";
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    15
by (stac Nat_unfold 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    16
by (rtac (singletonI RS UnI1) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    17
qed "Zero_RepI";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    18
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5187
diff changeset
    19
Goal "i: Nat ==> Suc_Rep(i) : Nat";
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    20
by (stac Nat_unfold 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    21
by (rtac (imageI RS UnI2) 1);
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5187
diff changeset
    22
by (assume_tac 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    23
qed "Suc_RepI";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    24
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    25
(*** Induction ***)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    26
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5187
diff changeset
    27
val major::prems = Goal
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    28
    "[| i: Nat;  P(Zero_Rep);   \
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    29
\       !!j. [| j: Nat; P(j) |] ==> P(Suc_Rep(j)) |]  ==> P(i)";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    30
by (rtac ([Nat_def, Nat_fun_mono, major] MRS def_induct) 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4032
diff changeset
    31
by (blast_tac (claset() addIs prems) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    32
qed "Nat_induct";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    33
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5187
diff changeset
    34
val prems = Goalw [Zero_def,Suc_def]
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    35
    "[| P(0);   \
3040
7d48671753da Introduced a generic "induct_tac" which picks up the right induction scheme
nipkow
parents: 3023
diff changeset
    36
\       !!n. P(n) ==> P(Suc(n)) |]  ==> P(n)";
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    37
by (rtac (Rep_Nat_inverse RS subst) 1);   (*types force good instantiation*)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    38
by (rtac (Rep_Nat RS Nat_induct) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    39
by (REPEAT (ares_tac prems 1
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    40
     ORELSE eresolve_tac [Abs_Nat_inverse RS subst] 1));
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    41
qed "nat_induct";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    42
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    43
(*Perform induction on n. *)
5187
55f07169cf5f Removed nat_case, nat_rec, and natE (now provided by datatype
berghofe
parents: 5148
diff changeset
    44
fun nat_ind_tac a i = 
55f07169cf5f Removed nat_case, nat_rec, and natE (now provided by datatype
berghofe
parents: 5148
diff changeset
    45
  res_inst_tac [("n",a)] nat_induct i  THEN  rename_last_tac a [""] (i+1);
3040
7d48671753da Introduced a generic "induct_tac" which picks up the right induction scheme
nipkow
parents: 3023
diff changeset
    46
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    47
(*A special form of induction for reasoning about m<n and m-n*)
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5187
diff changeset
    48
val prems = Goal
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    49
    "[| !!x. P x 0;  \
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    50
\       !!y. P 0 (Suc y);  \
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    51
\       !!x y. [| P x y |] ==> P (Suc x) (Suc y)  \
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    52
\    |] ==> P m n";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    53
by (res_inst_tac [("x","m")] spec 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    54
by (nat_ind_tac "n" 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    55
by (rtac allI 2);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    56
by (nat_ind_tac "x" 2);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    57
by (REPEAT (ares_tac (prems@[allI]) 1 ORELSE etac spec 1));
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    58
qed "diff_induct";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    59
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    60
(*** Isomorphisms: Abs_Nat and Rep_Nat ***)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    61
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    62
(*We can't take these properties as axioms, or take Abs_Nat==Inv(Rep_Nat),
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    63
  since we assume the isomorphism equations will one day be given by Isabelle*)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    64
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
    65
Goal "inj(Rep_Nat)";
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    66
by (rtac inj_inverseI 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    67
by (rtac Rep_Nat_inverse 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    68
qed "inj_Rep_Nat";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    69
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
    70
Goal "inj_on Abs_Nat Nat";
4830
bd73675adbed Added a few lemmas.
nipkow
parents: 4821
diff changeset
    71
by (rtac inj_on_inverseI 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    72
by (etac Abs_Nat_inverse 1);
4830
bd73675adbed Added a few lemmas.
nipkow
parents: 4821
diff changeset
    73
qed "inj_on_Abs_Nat";
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    74
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    75
(*** Distinctness of constructors ***)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    76
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
    77
Goalw [Zero_def,Suc_def] "Suc(m) ~= 0";
4830
bd73675adbed Added a few lemmas.
nipkow
parents: 4821
diff changeset
    78
by (rtac (inj_on_Abs_Nat RS inj_on_contraD) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    79
by (rtac Suc_Rep_not_Zero_Rep 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    80
by (REPEAT (resolve_tac [Rep_Nat, Suc_RepI, Zero_RepI] 1));
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    81
qed "Suc_not_Zero";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    82
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    83
bind_thm ("Zero_not_Suc", Suc_not_Zero RS not_sym);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    84
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    85
AddIffs [Suc_not_Zero,Zero_not_Suc];
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    86
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    87
bind_thm ("Suc_neq_Zero", (Suc_not_Zero RS notE));
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    88
val Zero_neq_Suc = sym RS Suc_neq_Zero;
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    89
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    90
(** Injectiveness of Suc **)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    91
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
    92
Goalw [Suc_def] "inj(Suc)";
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    93
by (rtac injI 1);
4830
bd73675adbed Added a few lemmas.
nipkow
parents: 4821
diff changeset
    94
by (dtac (inj_on_Abs_Nat RS inj_onD) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    95
by (REPEAT (resolve_tac [Rep_Nat, Suc_RepI] 1));
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    96
by (dtac (inj_Suc_Rep RS injD) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    97
by (etac (inj_Rep_Nat RS injD) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    98
qed "inj_Suc";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    99
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   100
val Suc_inject = inj_Suc RS injD;
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   101
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   102
Goal "(Suc(m)=Suc(n)) = (m=n)";
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   103
by (EVERY1 [rtac iffI, etac Suc_inject, etac arg_cong]); 
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   104
qed "Suc_Suc_eq";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   105
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   106
AddIffs [Suc_Suc_eq];
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   107
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   108
Goal "n ~= Suc(n)";
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   109
by (nat_ind_tac "n" 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   110
by (ALLGOALS Asm_simp_tac);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   111
qed "n_not_Suc_n";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   112
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   113
bind_thm ("Suc_n_not_n", n_not_Suc_n RS not_sym);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   114
5187
55f07169cf5f Removed nat_case, nat_rec, and natE (now provided by datatype
berghofe
parents: 5148
diff changeset
   115
(*** Basic properties of "less than" ***)
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   116
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   117
Goalw [wf_def, pred_nat_def] "wf(pred_nat)";
3718
d78cf498a88c Minor tidying to use Clarify_tac, etc.
paulson
parents: 3563
diff changeset
   118
by (Clarify_tac 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   119
by (nat_ind_tac "x" 1);
3236
882e125ed7da New pattern-matching definition of pred_nat
paulson
parents: 3143
diff changeset
   120
by (ALLGOALS Blast_tac);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   121
qed "wf_pred_nat";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   122
3378
11f4884a071a Moved "less_eq" to NatDef from Arith
paulson
parents: 3355
diff changeset
   123
(*Used in TFL/post.sml*)
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   124
Goalw [less_def] "(m,n) : pred_nat^+ = (m<n)";
3378
11f4884a071a Moved "less_eq" to NatDef from Arith
paulson
parents: 3355
diff changeset
   125
by (rtac refl 1);
11f4884a071a Moved "less_eq" to NatDef from Arith
paulson
parents: 3355
diff changeset
   126
qed "less_eq";
11f4884a071a Moved "less_eq" to NatDef from Arith
paulson
parents: 3355
diff changeset
   127
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   128
(** Introduction properties **)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   129
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5187
diff changeset
   130
Goalw [less_def] "[| i<j;  j<k |] ==> i<(k::nat)";
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   131
by (rtac (trans_trancl RS transD) 1);
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5187
diff changeset
   132
by (assume_tac 1);
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5187
diff changeset
   133
by (assume_tac 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   134
qed "less_trans";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   135
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   136
Goalw [less_def, pred_nat_def] "n < Suc(n)";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4032
diff changeset
   137
by (simp_tac (simpset() addsimps [r_into_trancl]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   138
qed "lessI";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   139
AddIffs [lessI];
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   140
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   141
(* i<j ==> i<Suc(j) *)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   142
bind_thm("less_SucI", lessI RSN (2, less_trans));
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   143
Addsimps [less_SucI];
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   144
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   145
Goal "0 < Suc(n)";
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   146
by (nat_ind_tac "n" 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   147
by (rtac lessI 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   148
by (etac less_trans 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   149
by (rtac lessI 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   150
qed "zero_less_Suc";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   151
AddIffs [zero_less_Suc];
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   152
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   153
(** Elimination properties **)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   154
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5187
diff changeset
   155
Goalw [less_def] "n<m ==> ~ m<(n::nat)";
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5187
diff changeset
   156
by (blast_tac (claset() addIs [wf_pred_nat, wf_trancl RS wf_asym])1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   157
qed "less_not_sym";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   158
5474
a2109bb8ce2b well-formed asym rules; also adds less_irrefl, le_refl since order_refl
paulson
parents: 5354
diff changeset
   159
(* [| n<m; ~P ==> m<n |] ==> P *)
a2109bb8ce2b well-formed asym rules; also adds less_irrefl, le_refl since order_refl
paulson
parents: 5354
diff changeset
   160
bind_thm ("less_asym", less_not_sym RS swap);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   161
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   162
Goalw [less_def] "~ n<(n::nat)";
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   163
by (rtac notI 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   164
by (etac (wf_pred_nat RS wf_trancl RS wf_irrefl) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   165
qed "less_not_refl";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   166
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   167
(* n<n ==> R *)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   168
bind_thm ("less_irrefl", (less_not_refl RS notE));
5474
a2109bb8ce2b well-formed asym rules; also adds less_irrefl, le_refl since order_refl
paulson
parents: 5354
diff changeset
   169
AddSEs [less_irrefl];
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   170
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5132
diff changeset
   171
Goal "n<m ==> m ~= (n::nat)";
5474
a2109bb8ce2b well-formed asym rules; also adds less_irrefl, le_refl since order_refl
paulson
parents: 5354
diff changeset
   172
by (Blast_tac 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   173
qed "less_not_refl2";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   174
5354
da63d9b35caf new theorems; adds [le_refl, less_imp_le] as simprules
paulson
parents: 5316
diff changeset
   175
(* s < t ==> s ~= t *)
da63d9b35caf new theorems; adds [le_refl, less_imp_le] as simprules
paulson
parents: 5316
diff changeset
   176
bind_thm ("less_not_refl3", less_not_refl2 RS not_sym);
da63d9b35caf new theorems; adds [le_refl, less_imp_le] as simprules
paulson
parents: 5316
diff changeset
   177
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   178
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5187
diff changeset
   179
val major::prems = Goalw [less_def, pred_nat_def]
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   180
    "[| i<k;  k=Suc(i) ==> P;  !!j. [| i<j;  k=Suc(j) |] ==> P \
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   181
\    |] ==> P";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   182
by (rtac (major RS tranclE) 1);
3236
882e125ed7da New pattern-matching definition of pred_nat
paulson
parents: 3143
diff changeset
   183
by (ALLGOALS Full_simp_tac); 
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   184
by (REPEAT_FIRST (bound_hyp_subst_tac ORELSE'
3236
882e125ed7da New pattern-matching definition of pred_nat
paulson
parents: 3143
diff changeset
   185
                  eresolve_tac (prems@[asm_rl, Pair_inject])));
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   186
qed "lessE";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   187
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   188
Goal "~ n<0";
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   189
by (rtac notI 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   190
by (etac lessE 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   191
by (etac Zero_neq_Suc 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   192
by (etac Zero_neq_Suc 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   193
qed "not_less0";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   194
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   195
AddIffs [not_less0];
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   196
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   197
(* n<0 ==> R *)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   198
bind_thm ("less_zeroE", not_less0 RS notE);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   199
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5187
diff changeset
   200
val [major,less,eq] = Goal
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   201
    "[| m < Suc(n);  m<n ==> P;  m=n ==> P |] ==> P";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   202
by (rtac (major RS lessE) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   203
by (rtac eq 1);
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   204
by (Blast_tac 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   205
by (rtac less 1);
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   206
by (Blast_tac 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   207
qed "less_SucE";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   208
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   209
Goal "(m < Suc(n)) = (m < n | m = n)";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4032
diff changeset
   210
by (blast_tac (claset() addSEs [less_SucE] addIs [less_trans]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   211
qed "less_Suc_eq";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   212
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   213
Goal "(n<1) = (n=0)";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4032
diff changeset
   214
by (simp_tac (simpset() addsimps [less_Suc_eq]) 1);
3484
1e93eb09ebb9 Added the following lemmas tp Divides and a few others to Arith and NatDef:
nipkow
parents: 3457
diff changeset
   215
qed "less_one";
1e93eb09ebb9 Added the following lemmas tp Divides and a few others to Arith and NatDef:
nipkow
parents: 3457
diff changeset
   216
AddIffs [less_one];
1e93eb09ebb9 Added the following lemmas tp Divides and a few others to Arith and NatDef:
nipkow
parents: 3457
diff changeset
   217
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5132
diff changeset
   218
Goal "m<n ==> Suc(m) < Suc(n)";
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   219
by (etac rev_mp 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   220
by (nat_ind_tac "n" 1);
5474
a2109bb8ce2b well-formed asym rules; also adds less_irrefl, le_refl since order_refl
paulson
parents: 5354
diff changeset
   221
by (ALLGOALS (fast_tac (claset() addEs [less_trans, lessE])));
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   222
qed "Suc_mono";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   223
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   224
(*"Less than" is a linear ordering*)
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   225
Goal "m<n | m=n | n<(m::nat)";
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   226
by (nat_ind_tac "m" 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   227
by (nat_ind_tac "n" 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   228
by (rtac (refl RS disjI1 RS disjI2) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   229
by (rtac (zero_less_Suc RS disjI1) 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4032
diff changeset
   230
by (blast_tac (claset() addIs [Suc_mono, less_SucI] addEs [lessE]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   231
qed "less_linear";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   232
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   233
Goal "!!m::nat. (m ~= n) = (m<n | n<m)";
4737
4544290d5a6b The theorem nat_neqE, and some tidying
paulson
parents: 4686
diff changeset
   234
by (cut_facts_tac [less_linear] 1);
5500
7e0ed3e31590 new theorem less_Suc_eq_le and le_simps
paulson
parents: 5478
diff changeset
   235
by (Blast_tac 1);
4737
4544290d5a6b The theorem nat_neqE, and some tidying
paulson
parents: 4686
diff changeset
   236
qed "nat_neq_iff";
4544290d5a6b The theorem nat_neqE, and some tidying
paulson
parents: 4686
diff changeset
   237
7030
53934985426a getting rid of qed_goal
paulson
parents: 6115
diff changeset
   238
val [major,eqCase,lessCase] = Goal 
53934985426a getting rid of qed_goal
paulson
parents: 6115
diff changeset
   239
   "[| (m::nat)<n ==> P n m; m=n ==> P n m; n<m ==> P n m |] ==> P n m";
53934985426a getting rid of qed_goal
paulson
parents: 6115
diff changeset
   240
by (rtac (less_linear RS disjE) 1);
53934985426a getting rid of qed_goal
paulson
parents: 6115
diff changeset
   241
by (etac disjE 2);
53934985426a getting rid of qed_goal
paulson
parents: 6115
diff changeset
   242
by (etac lessCase 1);
53934985426a getting rid of qed_goal
paulson
parents: 6115
diff changeset
   243
by (etac (sym RS eqCase) 1);
53934985426a getting rid of qed_goal
paulson
parents: 6115
diff changeset
   244
by (etac major 1);
53934985426a getting rid of qed_goal
paulson
parents: 6115
diff changeset
   245
qed "nat_less_cases";
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   246
4745
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   247
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   248
(** Inductive (?) properties **)
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   249
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5132
diff changeset
   250
Goal "[| m<n; Suc m ~= n |] ==> Suc(m) < n";
4745
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   251
by (full_simp_tac (simpset() addsimps [nat_neq_iff]) 1);
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   252
by (blast_tac (claset() addSEs [less_irrefl, less_SucE] addEs [less_asym]) 1);
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   253
qed "Suc_lessI";
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   254
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5187
diff changeset
   255
Goal "Suc(m) < n ==> m<n";
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5187
diff changeset
   256
by (etac rev_mp 1);
4745
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   257
by (nat_ind_tac "n" 1);
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   258
by (ALLGOALS (fast_tac (claset() addSIs [lessI RS less_SucI]
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   259
                                 addEs  [less_trans, lessE])));
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   260
qed "Suc_lessD";
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   261
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5187
diff changeset
   262
val [major,minor] = Goal 
4745
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   263
    "[| Suc(i)<k;  !!j. [| i<j;  k=Suc(j) |] ==> P \
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   264
\    |] ==> P";
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   265
by (rtac (major RS lessE) 1);
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   266
by (etac (lessI RS minor) 1);
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   267
by (etac (Suc_lessD RS minor) 1);
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   268
by (assume_tac 1);
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   269
qed "Suc_lessE";
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   270
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5132
diff changeset
   271
Goal "Suc(m) < Suc(n) ==> m<n";
4745
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   272
by (blast_tac (claset() addEs [lessE, make_elim Suc_lessD]) 1);
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   273
qed "Suc_less_SucD";
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   274
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   275
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   276
Goal "(Suc(m) < Suc(n)) = (m<n)";
4745
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   277
by (EVERY1 [rtac iffI, etac Suc_less_SucD, etac Suc_mono]);
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   278
qed "Suc_less_eq";
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   279
Addsimps [Suc_less_eq];
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   280
6109
82b50115564c Simplified arithmetic.
nipkow
parents: 6075
diff changeset
   281
(*Goal "~(Suc(n) < n)";
4745
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   282
by (blast_tac (claset() addEs [Suc_lessD RS less_irrefl]) 1);
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   283
qed "not_Suc_n_less_n";
6109
82b50115564c Simplified arithmetic.
nipkow
parents: 6075
diff changeset
   284
Addsimps [not_Suc_n_less_n];*)
4745
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   285
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5132
diff changeset
   286
Goal "i<j ==> j<k --> Suc i < k";
4745
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   287
by (nat_ind_tac "k" 1);
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   288
by (ALLGOALS (asm_simp_tac (simpset())));
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   289
by (asm_simp_tac (simpset() addsimps [less_Suc_eq]) 1);
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   290
by (blast_tac (claset() addDs [Suc_lessD]) 1);
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   291
qed_spec_mp "less_trans_Suc";
b855a7094195 re-ordered proofs
paulson
parents: 4737
diff changeset
   292
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   293
(*Can be used with less_Suc_eq to get n=m | n<m *)
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   294
Goal "(~ m < n) = (n < Suc(m))";
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   295
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   296
by (ALLGOALS Asm_simp_tac);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   297
qed "not_less_eq";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   298
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   299
(*Complete induction, aka course-of-values induction*)
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5187
diff changeset
   300
val prems = Goalw [less_def]
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   301
    "[| !!n. [| ! m::nat. m<n --> P(m) |] ==> P(n) |]  ==>  P(n)";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   302
by (wf_ind_tac "n" [wf_pred_nat RS wf_trancl] 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   303
by (eresolve_tac prems 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   304
qed "less_induct";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   305
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   306
(*** Properties of <= ***)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   307
5500
7e0ed3e31590 new theorem less_Suc_eq_le and le_simps
paulson
parents: 5478
diff changeset
   308
(*Was le_eq_less_Suc, but this orientation is more useful*)
7e0ed3e31590 new theorem less_Suc_eq_le and le_simps
paulson
parents: 5478
diff changeset
   309
Goalw [le_def] "(m < Suc n) = (m <= n)";
7e0ed3e31590 new theorem less_Suc_eq_le and le_simps
paulson
parents: 5478
diff changeset
   310
by (rtac (not_less_eq RS sym) 1);
7e0ed3e31590 new theorem less_Suc_eq_le and le_simps
paulson
parents: 5478
diff changeset
   311
qed "less_Suc_eq_le";
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   312
3343
45986997f1fe Renamed lessD to Suc_leI
paulson
parents: 3308
diff changeset
   313
(*  m<=n ==> m < Suc n  *)
5500
7e0ed3e31590 new theorem less_Suc_eq_le and le_simps
paulson
parents: 5478
diff changeset
   314
bind_thm ("le_imp_less_Suc", less_Suc_eq_le RS iffD2);
3343
45986997f1fe Renamed lessD to Suc_leI
paulson
parents: 3308
diff changeset
   315
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   316
Goalw [le_def] "0 <= n";
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   317
by (rtac not_less0 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   318
qed "le0";
6075
c148037f53c6 Remoaved a few now redundant rewrite rules.
nipkow
parents: 5983
diff changeset
   319
AddIffs [le0];
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   320
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   321
Goalw [le_def] "~ Suc n <= n";
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   322
by (Simp_tac 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   323
qed "Suc_n_not_le_n";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   324
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   325
Goalw [le_def] "(i <= 0) = (i = 0)";
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   326
by (nat_ind_tac "i" 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   327
by (ALLGOALS Asm_simp_tac);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   328
qed "le_0_eq";
4614
122015efd4e1 New AddIffs le_0_eq and neq0_conv
paulson
parents: 4599
diff changeset
   329
AddIffs [le_0_eq];
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   330
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5132
diff changeset
   331
Goal "(m <= Suc(n)) = (m<=n | m = Suc n)";
5500
7e0ed3e31590 new theorem less_Suc_eq_le and le_simps
paulson
parents: 5478
diff changeset
   332
by (simp_tac (simpset() delsimps [less_Suc_eq_le]
7e0ed3e31590 new theorem less_Suc_eq_le and le_simps
paulson
parents: 5478
diff changeset
   333
			addsimps [less_Suc_eq_le RS sym, less_Suc_eq]) 1);
3355
0d955bcf8e0a New theorem le_Suc_eq
paulson
parents: 3343
diff changeset
   334
qed "le_Suc_eq";
0d955bcf8e0a New theorem le_Suc_eq
paulson
parents: 3343
diff changeset
   335
4614
122015efd4e1 New AddIffs le_0_eq and neq0_conv
paulson
parents: 4599
diff changeset
   336
(* [| m <= Suc n;  m <= n ==> R;  m = Suc n ==> R |] ==> R *)
122015efd4e1 New AddIffs le_0_eq and neq0_conv
paulson
parents: 4599
diff changeset
   337
bind_thm ("le_SucE", le_Suc_eq RS iffD1 RS disjE);
122015efd4e1 New AddIffs le_0_eq and neq0_conv
paulson
parents: 4599
diff changeset
   338
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5187
diff changeset
   339
Goalw [le_def] "~n<m ==> m<=(n::nat)";
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5187
diff changeset
   340
by (assume_tac 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   341
qed "leI";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   342
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5187
diff changeset
   343
Goalw [le_def] "m<=n ==> ~ n < (m::nat)";
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5187
diff changeset
   344
by (assume_tac 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   345
qed "leD";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   346
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   347
val leE = make_elim leD;
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   348
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   349
Goal "(~n<m) = (m<=(n::nat))";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4032
diff changeset
   350
by (blast_tac (claset() addIs [leI] addEs [leE]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   351
qed "not_less_iff_le";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   352
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5132
diff changeset
   353
Goalw [le_def] "~ m <= n ==> n<(m::nat)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   354
by (Blast_tac 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   355
qed "not_leE";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   356
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   357
Goalw [le_def] "(~n<=m) = (m<(n::nat))";
4599
3a4348a3d6ff New max, min theorems
paulson
parents: 4535
diff changeset
   358
by (Simp_tac 1);
3a4348a3d6ff New max, min theorems
paulson
parents: 4535
diff changeset
   359
qed "not_le_iff_less";
3a4348a3d6ff New max, min theorems
paulson
parents: 4535
diff changeset
   360
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5132
diff changeset
   361
Goalw [le_def] "m < n ==> Suc(m) <= n";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4032
diff changeset
   362
by (simp_tac (simpset() addsimps [less_Suc_eq]) 1);
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4032
diff changeset
   363
by (blast_tac (claset() addSEs [less_irrefl,less_asym]) 1);
3343
45986997f1fe Renamed lessD to Suc_leI
paulson
parents: 3308
diff changeset
   364
qed "Suc_leI";  (*formerly called lessD*)
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   365
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5132
diff changeset
   366
Goalw [le_def] "Suc(m) <= n ==> m <= n";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4032
diff changeset
   367
by (asm_full_simp_tac (simpset() addsimps [less_Suc_eq]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   368
qed "Suc_leD";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   369
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   370
(* stronger version of Suc_leD *)
5148
74919e8f221c More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5143
diff changeset
   371
Goalw [le_def] "Suc m <= n ==> m < n";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4032
diff changeset
   372
by (asm_full_simp_tac (simpset() addsimps [less_Suc_eq]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   373
by (cut_facts_tac [less_linear] 1);
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   374
by (Blast_tac 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   375
qed "Suc_le_lessD";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   376
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   377
Goal "(Suc m <= n) = (m < n)";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4032
diff changeset
   378
by (blast_tac (claset() addIs [Suc_leI, Suc_le_lessD]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   379
qed "Suc_le_eq";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   380
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5132
diff changeset
   381
Goalw [le_def] "m <= n ==> m <= Suc n";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4032
diff changeset
   382
by (blast_tac (claset() addDs [Suc_lessD]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   383
qed "le_SucI";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   384
Addsimps[le_SucI];
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   385
6109
82b50115564c Simplified arithmetic.
nipkow
parents: 6075
diff changeset
   386
(*bind_thm ("le_Suc", not_Suc_n_less_n RS leI);*)
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   387
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5132
diff changeset
   388
Goalw [le_def] "m < n ==> m <= (n::nat)";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4032
diff changeset
   389
by (blast_tac (claset() addEs [less_asym]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   390
qed "less_imp_le";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   391
5591
paulson
parents: 5500
diff changeset
   392
(*For instance, (Suc m < Suc n)  =   (Suc m <= n)  =  (m<n) *)
paulson
parents: 5500
diff changeset
   393
val le_simps = [less_imp_le, less_Suc_eq_le, Suc_le_eq];
paulson
parents: 5500
diff changeset
   394
5354
da63d9b35caf new theorems; adds [le_refl, less_imp_le] as simprules
paulson
parents: 5316
diff changeset
   395
3343
45986997f1fe Renamed lessD to Suc_leI
paulson
parents: 3308
diff changeset
   396
(** Equivalence of m<=n and  m<n | m=n **)
45986997f1fe Renamed lessD to Suc_leI
paulson
parents: 3308
diff changeset
   397
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5132
diff changeset
   398
Goalw [le_def] "m <= n ==> m < n | m=(n::nat)";
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   399
by (cut_facts_tac [less_linear] 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4032
diff changeset
   400
by (blast_tac (claset() addEs [less_irrefl,less_asym]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   401
qed "le_imp_less_or_eq";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   402
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5132
diff changeset
   403
Goalw [le_def] "m<n | m=n ==> m <=(n::nat)";
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   404
by (cut_facts_tac [less_linear] 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4032
diff changeset
   405
by (blast_tac (claset() addSEs [less_irrefl] addEs [less_asym]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   406
qed "less_or_eq_imp_le";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   407
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   408
Goal "(m <= (n::nat)) = (m < n | m=n)";
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   409
by (REPEAT(ares_tac [iffI,less_or_eq_imp_le,le_imp_less_or_eq] 1));
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   410
qed "le_eq_less_or_eq";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   411
4635
c448e09d0fca New theorem eq_imp_le
paulson
parents: 4614
diff changeset
   412
(*Useful with Blast_tac.   m=n ==> m<=n *)
c448e09d0fca New theorem eq_imp_le
paulson
parents: 4614
diff changeset
   413
bind_thm ("eq_imp_le", disjI2 RS less_or_eq_imp_le);
c448e09d0fca New theorem eq_imp_le
paulson
parents: 4614
diff changeset
   414
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   415
Goal "n <= (n::nat)";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4032
diff changeset
   416
by (simp_tac (simpset() addsimps [le_eq_less_or_eq]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   417
qed "le_refl";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   418
5354
da63d9b35caf new theorems; adds [le_refl, less_imp_le] as simprules
paulson
parents: 5316
diff changeset
   419
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5132
diff changeset
   420
Goal "[| i <= j; j < k |] ==> i < (k::nat)";
4468
paulson
parents: 4423
diff changeset
   421
by (blast_tac (claset() addSDs [le_imp_less_or_eq]
paulson
parents: 4423
diff changeset
   422
	                addIs [less_trans]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   423
qed "le_less_trans";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   424
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5132
diff changeset
   425
Goal "[| i < j; j <= k |] ==> i < (k::nat)";
4468
paulson
parents: 4423
diff changeset
   426
by (blast_tac (claset() addSDs [le_imp_less_or_eq]
paulson
parents: 4423
diff changeset
   427
	                addIs [less_trans]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   428
qed "less_le_trans";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   429
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5132
diff changeset
   430
Goal "[| i <= j; j <= k |] ==> i <= (k::nat)";
4468
paulson
parents: 4423
diff changeset
   431
by (blast_tac (claset() addSDs [le_imp_less_or_eq]
paulson
parents: 4423
diff changeset
   432
	                addIs [less_or_eq_imp_le, less_trans]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   433
qed "le_trans";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   434
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5132
diff changeset
   435
Goal "[| m <= n; n <= m |] ==> m = (n::nat)";
4468
paulson
parents: 4423
diff changeset
   436
(*order_less_irrefl could make this proof fail*)
paulson
parents: 4423
diff changeset
   437
by (blast_tac (claset() addSDs [le_imp_less_or_eq]
paulson
parents: 4423
diff changeset
   438
	                addSEs [less_irrefl] addEs [less_asym]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   439
qed "le_anti_sym";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   440
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   441
Goal "(Suc(n) <= Suc(m)) = (n <= m)";
5500
7e0ed3e31590 new theorem less_Suc_eq_le and le_simps
paulson
parents: 5478
diff changeset
   442
by (simp_tac (simpset() addsimps le_simps) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   443
qed "Suc_le_mono";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   444
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   445
AddIffs [Suc_le_mono];
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   446
5500
7e0ed3e31590 new theorem less_Suc_eq_le and le_simps
paulson
parents: 5478
diff changeset
   447
(* Axiom 'order_less_le' of class 'order': *)
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   448
Goal "(m::nat) < n = (m <= n & m ~= n)";
4737
4544290d5a6b The theorem nat_neqE, and some tidying
paulson
parents: 4686
diff changeset
   449
by (simp_tac (simpset() addsimps [le_def, nat_neq_iff]) 1);
4544290d5a6b The theorem nat_neqE, and some tidying
paulson
parents: 4686
diff changeset
   450
by (blast_tac (claset() addSEs [less_asym]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   451
qed "nat_less_le";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   452
5354
da63d9b35caf new theorems; adds [le_refl, less_imp_le] as simprules
paulson
parents: 5316
diff changeset
   453
(* [| m <= n; m ~= n |] ==> m < n *)
da63d9b35caf new theorems; adds [le_refl, less_imp_le] as simprules
paulson
parents: 5316
diff changeset
   454
bind_thm ("le_neq_implies_less", [nat_less_le, conjI] MRS iffD2);
da63d9b35caf new theorems; adds [le_refl, less_imp_le] as simprules
paulson
parents: 5316
diff changeset
   455
4640
ac6cf9f18653 Congruence rules use == in premises now.
nipkow
parents: 4635
diff changeset
   456
(* Axiom 'linorder_linear' of class 'linorder': *)
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   457
Goal "(m::nat) <= n | n <= m";
4640
ac6cf9f18653 Congruence rules use == in premises now.
nipkow
parents: 4635
diff changeset
   458
by (simp_tac (simpset() addsimps [le_eq_less_or_eq]) 1);
ac6cf9f18653 Congruence rules use == in premises now.
nipkow
parents: 4635
diff changeset
   459
by (cut_facts_tac [less_linear] 1);
5132
24f992a25adc isatool expandshort;
wenzelm
parents: 5069
diff changeset
   460
by (Blast_tac 1);
4640
ac6cf9f18653 Congruence rules use == in premises now.
nipkow
parents: 4635
diff changeset
   461
qed "nat_le_linear";
ac6cf9f18653 Congruence rules use == in premises now.
nipkow
parents: 4635
diff changeset
   462
5354
da63d9b35caf new theorems; adds [le_refl, less_imp_le] as simprules
paulson
parents: 5316
diff changeset
   463
Goal "~ n < m ==> (n < Suc m) = (n = m)";
da63d9b35caf new theorems; adds [le_refl, less_imp_le] as simprules
paulson
parents: 5316
diff changeset
   464
by (blast_tac (claset() addSEs [less_SucE]) 1);
da63d9b35caf new theorems; adds [le_refl, less_imp_le] as simprules
paulson
parents: 5316
diff changeset
   465
qed "not_less_less_Suc_eq";
da63d9b35caf new theorems; adds [le_refl, less_imp_le] as simprules
paulson
parents: 5316
diff changeset
   466
da63d9b35caf new theorems; adds [le_refl, less_imp_le] as simprules
paulson
parents: 5316
diff changeset
   467
da63d9b35caf new theorems; adds [le_refl, less_imp_le] as simprules
paulson
parents: 5316
diff changeset
   468
(*Rewrite (n < Suc m) to (n=m) if  ~ n<m or m<=n hold.
da63d9b35caf new theorems; adds [le_refl, less_imp_le] as simprules
paulson
parents: 5316
diff changeset
   469
  Not suitable as default simprules because they often lead to looping*)
da63d9b35caf new theorems; adds [le_refl, less_imp_le] as simprules
paulson
parents: 5316
diff changeset
   470
val not_less_simps = [not_less_less_Suc_eq, leD RS not_less_less_Suc_eq];
4640
ac6cf9f18653 Congruence rules use == in premises now.
nipkow
parents: 4635
diff changeset
   471
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   472
(** LEAST -- the least number operator **)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   473
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   474
Goal "(! m::nat. P m --> n <= m) = (! m. m < n --> ~ P m)";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4032
diff changeset
   475
by (blast_tac (claset() addIs [leI] addEs [leE]) 1);
3143
d60e49b86c6a Modified def of Least, which, as Markus correctly complained, looked like
nipkow
parents: 3085
diff changeset
   476
val lemma = result();
d60e49b86c6a Modified def of Least, which, as Markus correctly complained, looked like
nipkow
parents: 3085
diff changeset
   477
d60e49b86c6a Modified def of Least, which, as Markus correctly complained, looked like
nipkow
parents: 3085
diff changeset
   478
(* This is an old def of Least for nat, which is derived for compatibility *)
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   479
Goalw [Least_def]
3143
d60e49b86c6a Modified def of Least, which, as Markus correctly complained, looked like
nipkow
parents: 3085
diff changeset
   480
  "(LEAST n::nat. P n) == (@n. P(n) & (ALL m. m < n --> ~P(m)))";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4032
diff changeset
   481
by (simp_tac (simpset() addsimps [lemma]) 1);
3143
d60e49b86c6a Modified def of Least, which, as Markus correctly complained, looked like
nipkow
parents: 3085
diff changeset
   482
qed "Least_nat_def";
d60e49b86c6a Modified def of Least, which, as Markus correctly complained, looked like
nipkow
parents: 3085
diff changeset
   483
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5187
diff changeset
   484
val [prem1,prem2] = Goalw [Least_nat_def]
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3768
diff changeset
   485
    "[| P(k::nat);  !!x. x<k ==> ~P(x) |] ==> (LEAST x. P(x)) = k";
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   486
by (rtac select_equality 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4032
diff changeset
   487
by (blast_tac (claset() addSIs [prem1,prem2]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   488
by (cut_facts_tac [less_linear] 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4032
diff changeset
   489
by (blast_tac (claset() addSIs [prem1] addSDs [prem2]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   490
qed "Least_equality";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   491
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5187
diff changeset
   492
Goal "P(k::nat) ==> P(LEAST x. P(x))";
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5187
diff changeset
   493
by (etac rev_mp 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   494
by (res_inst_tac [("n","k")] less_induct 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   495
by (rtac impI 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   496
by (rtac classical 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   497
by (res_inst_tac [("s","n")] (Least_equality RS ssubst) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   498
by (assume_tac 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   499
by (assume_tac 2);
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   500
by (Blast_tac 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   501
qed "LeastI";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   502
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   503
(*Proof is almost identical to the one above!*)
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5187
diff changeset
   504
Goal "P(k::nat) ==> (LEAST x. P(x)) <= k";
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5187
diff changeset
   505
by (etac rev_mp 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   506
by (res_inst_tac [("n","k")] less_induct 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   507
by (rtac impI 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   508
by (rtac classical 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   509
by (res_inst_tac [("s","n")] (Least_equality RS ssubst) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   510
by (assume_tac 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   511
by (rtac le_refl 2);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4032
diff changeset
   512
by (blast_tac (claset() addIs [less_imp_le,le_trans]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   513
qed "Least_le";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   514
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5187
diff changeset
   515
Goal "k < (LEAST x. P(x)) ==> ~P(k::nat)";
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   516
by (rtac notI 1);
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5187
diff changeset
   517
by (etac (rewrite_rule [le_def] Least_le RS notE) 1 THEN assume_tac 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   518
qed "not_less_Least";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   519
5983
79e301a6a51b At last: linear arithmetic for nat!
nipkow
parents: 5654
diff changeset
   520
(* [| m ~= n; m < n ==> P; n < m ==> P |] ==> P *)
4737
4544290d5a6b The theorem nat_neqE, and some tidying
paulson
parents: 4686
diff changeset
   521
bind_thm("nat_neqE", nat_neq_iff RS iffD1 RS disjE);
7064
b053e0ab9f60 New lemmas by Stefan Merz.
nipkow
parents: 7030
diff changeset
   522
b053e0ab9f60 New lemmas by Stefan Merz.
nipkow
parents: 7030
diff changeset
   523
Goal "(S::nat set) ~= {} ==> ? x:S. ! y:S. x <= y";
b053e0ab9f60 New lemmas by Stefan Merz.
nipkow
parents: 7030
diff changeset
   524
by (cut_facts_tac [wf_pred_nat RS wf_trancl RS (wf_eq_minimal RS iffD1)] 1);
b053e0ab9f60 New lemmas by Stefan Merz.
nipkow
parents: 7030
diff changeset
   525
by (dres_inst_tac [("x","S")] spec 1);
b053e0ab9f60 New lemmas by Stefan Merz.
nipkow
parents: 7030
diff changeset
   526
by (Asm_full_simp_tac 1);
b053e0ab9f60 New lemmas by Stefan Merz.
nipkow
parents: 7030
diff changeset
   527
by (etac impE 1);
b053e0ab9f60 New lemmas by Stefan Merz.
nipkow
parents: 7030
diff changeset
   528
by (Force_tac 1);
b053e0ab9f60 New lemmas by Stefan Merz.
nipkow
parents: 7030
diff changeset
   529
by (force_tac (claset(), simpset() addsimps [less_eq,not_le_iff_less]) 1);
b053e0ab9f60 New lemmas by Stefan Merz.
nipkow
parents: 7030
diff changeset
   530
qed "nonempty_has_least";