src/HOL/NatDef.ML
author paulson
Wed, 23 Jul 1997 11:50:26 +0200
changeset 3563 c4f13747489f
parent 3484 1e93eb09ebb9
child 3718 d78cf498a88c
permissions -rw-r--r--
Uses new version of Datatype.occs_in_prems
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
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
     7
goal thy "mono(%X. {Zero_Rep} Un (Suc_Rep``X))";
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*)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    14
goal thy "Zero_Rep: Nat";
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
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    19
val prems = goal thy "i: Nat ==> Suc_Rep(i) : Nat";
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);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    22
by (resolve_tac prems 1);
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
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    27
val major::prems = goal thy
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);
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
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
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    34
val prems = goalw thy [Zero_def,Suc_def]
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. *)
3563
c4f13747489f Uses new version of Datatype.occs_in_prems
paulson
parents: 3484
diff changeset
    44
local fun raw_nat_ind_tac a i = 
c4f13747489f Uses new version of Datatype.occs_in_prems
paulson
parents: 3484
diff changeset
    45
    res_inst_tac [("n",a)] nat_induct i  THEN  rename_last_tac a [""] (i+1)
c4f13747489f Uses new version of Datatype.occs_in_prems
paulson
parents: 3484
diff changeset
    46
in
c4f13747489f Uses new version of Datatype.occs_in_prems
paulson
parents: 3484
diff changeset
    47
val nat_ind_tac = Datatype.occs_in_prems raw_nat_ind_tac
c4f13747489f Uses new version of Datatype.occs_in_prems
paulson
parents: 3484
diff changeset
    48
end;
3040
7d48671753da Introduced a generic "induct_tac" which picks up the right induction scheme
nipkow
parents: 3023
diff changeset
    49
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    50
(*A special form of induction for reasoning about m<n and m-n*)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    51
val prems = goal thy
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    52
    "[| !!x. P x 0;  \
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    53
\       !!y. P 0 (Suc y);  \
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    54
\       !!x y. [| P x y |] ==> P (Suc x) (Suc y)  \
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    55
\    |] ==> P m n";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    56
by (res_inst_tac [("x","m")] spec 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    57
by (nat_ind_tac "n" 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    58
by (rtac allI 2);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    59
by (nat_ind_tac "x" 2);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    60
by (REPEAT (ares_tac (prems@[allI]) 1 ORELSE etac spec 1));
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    61
qed "diff_induct";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    62
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    63
(*Case analysis on the natural numbers*)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    64
val prems = goal thy 
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    65
    "[| n=0 ==> P;  !!x. n = Suc(x) ==> P |] ==> P";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    66
by (subgoal_tac "n=0 | (EX x. n = Suc(x))" 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    67
by (fast_tac (!claset addSEs prems) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    68
by (nat_ind_tac "n" 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    69
by (rtac (refl RS disjI1) 1);
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
    70
by (Blast_tac 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    71
qed "natE";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    72
3282
c31e6239d4c9 Added exhaustion thm and exhaust_tac for each datatype.
nipkow
parents: 3236
diff changeset
    73
(*Install 'automatic' induction tactic, pretending nat is a datatype *)
3292
8b143c196d42 Added rotation to exhaust_tac.
nipkow
parents: 3282
diff changeset
    74
(*except for induct_tac and exhaust_tac, everything is dummy*)
3282
c31e6239d4c9 Added exhaustion thm and exhaust_tac for each datatype.
nipkow
parents: 3236
diff changeset
    75
datatypes := [("nat",{case_const = Bound 0, case_rewrites = [],
c31e6239d4c9 Added exhaustion thm and exhaust_tac for each datatype.
nipkow
parents: 3236
diff changeset
    76
  constructors = [], induct_tac = nat_ind_tac,
3292
8b143c196d42 Added rotation to exhaust_tac.
nipkow
parents: 3282
diff changeset
    77
  exhaustion = natE,
8b143c196d42 Added rotation to exhaust_tac.
nipkow
parents: 3282
diff changeset
    78
  exhaust_tac = fn v => ALLNEWSUBGOALS (res_inst_tac [("n",v)] natE)
8b143c196d42 Added rotation to exhaust_tac.
nipkow
parents: 3282
diff changeset
    79
                                       (rotate_tac ~1),
3282
c31e6239d4c9 Added exhaustion thm and exhaust_tac for each datatype.
nipkow
parents: 3236
diff changeset
    80
  nchotomy = flexpair_def, case_cong = flexpair_def})];
c31e6239d4c9 Added exhaustion thm and exhaust_tac for each datatype.
nipkow
parents: 3236
diff changeset
    81
c31e6239d4c9 Added exhaustion thm and exhaust_tac for each datatype.
nipkow
parents: 3236
diff changeset
    82
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    83
(*** Isomorphisms: Abs_Nat and Rep_Nat ***)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    84
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    85
(*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
    86
  since we assume the isomorphism equations will one day be given by Isabelle*)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    87
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    88
goal thy "inj(Rep_Nat)";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    89
by (rtac inj_inverseI 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    90
by (rtac Rep_Nat_inverse 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    91
qed "inj_Rep_Nat";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    92
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    93
goal thy "inj_onto Abs_Nat Nat";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    94
by (rtac inj_onto_inverseI 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    95
by (etac Abs_Nat_inverse 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    96
qed "inj_onto_Abs_Nat";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    97
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    98
(*** Distinctness of constructors ***)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
    99
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   100
goalw thy [Zero_def,Suc_def] "Suc(m) ~= 0";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   101
by (rtac (inj_onto_Abs_Nat RS inj_onto_contraD) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   102
by (rtac Suc_Rep_not_Zero_Rep 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   103
by (REPEAT (resolve_tac [Rep_Nat, Suc_RepI, Zero_RepI] 1));
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   104
qed "Suc_not_Zero";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   105
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   106
bind_thm ("Zero_not_Suc", Suc_not_Zero RS not_sym);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   107
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   108
AddIffs [Suc_not_Zero,Zero_not_Suc];
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   109
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   110
bind_thm ("Suc_neq_Zero", (Suc_not_Zero RS notE));
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   111
val Zero_neq_Suc = sym RS Suc_neq_Zero;
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   112
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   113
(** Injectiveness of Suc **)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   114
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   115
goalw thy [Suc_def] "inj(Suc)";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   116
by (rtac injI 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   117
by (dtac (inj_onto_Abs_Nat RS inj_ontoD) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   118
by (REPEAT (resolve_tac [Rep_Nat, Suc_RepI] 1));
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   119
by (dtac (inj_Suc_Rep RS injD) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   120
by (etac (inj_Rep_Nat RS injD) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   121
qed "inj_Suc";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   122
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   123
val Suc_inject = inj_Suc RS injD;
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   124
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   125
goal thy "(Suc(m)=Suc(n)) = (m=n)";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   126
by (EVERY1 [rtac iffI, etac Suc_inject, etac arg_cong]); 
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   127
qed "Suc_Suc_eq";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   128
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   129
AddIffs [Suc_Suc_eq];
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   130
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   131
goal thy "n ~= Suc(n)";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   132
by (nat_ind_tac "n" 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   133
by (ALLGOALS Asm_simp_tac);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   134
qed "n_not_Suc_n";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   135
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   136
bind_thm ("Suc_n_not_n", n_not_Suc_n RS not_sym);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   137
3236
882e125ed7da New pattern-matching definition of pred_nat
paulson
parents: 3143
diff changeset
   138
goal thy "!!n. n ~= 0 ==> EX m. n = Suc m";
3457
a8ab7c64817c Ran expandshort
paulson
parents: 3378
diff changeset
   139
by (rtac natE 1);
3236
882e125ed7da New pattern-matching definition of pred_nat
paulson
parents: 3143
diff changeset
   140
by (REPEAT (Blast_tac 1));
882e125ed7da New pattern-matching definition of pred_nat
paulson
parents: 3143
diff changeset
   141
qed "not0_implies_Suc";
882e125ed7da New pattern-matching definition of pred_nat
paulson
parents: 3143
diff changeset
   142
882e125ed7da New pattern-matching definition of pred_nat
paulson
parents: 3143
diff changeset
   143
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   144
(*** nat_case -- the selection operator for nat ***)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   145
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   146
goalw thy [nat_case_def] "nat_case a f 0 = a";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   147
by (blast_tac (!claset addIs [select_equality]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   148
qed "nat_case_0";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   149
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   150
goalw thy [nat_case_def] "nat_case a f (Suc k) = f(k)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   151
by (blast_tac (!claset addIs [select_equality]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   152
qed "nat_case_Suc";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   153
3236
882e125ed7da New pattern-matching definition of pred_nat
paulson
parents: 3143
diff changeset
   154
goalw thy [wf_def, pred_nat_def] "wf(pred_nat)";
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   155
by (strip_tac 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   156
by (nat_ind_tac "x" 1);
3236
882e125ed7da New pattern-matching definition of pred_nat
paulson
parents: 3143
diff changeset
   157
by (ALLGOALS Blast_tac);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   158
qed "wf_pred_nat";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   159
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   160
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   161
(*** nat_rec -- by wf recursion on pred_nat ***)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   162
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   163
(* The unrolling rule for nat_rec *)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   164
goal thy
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   165
   "(%n. nat_rec c d n) = wfrec pred_nat (%f. nat_case ?c (%m. ?d m (f m)))";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   166
  by (simp_tac (HOL_ss addsimps [nat_rec_def]) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   167
bind_thm("nat_rec_unfold", wf_pred_nat RS 
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   168
                            ((result() RS eq_reflection) RS def_wfrec));
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   169
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   170
(*---------------------------------------------------------------------------
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   171
 * Old:
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   172
 * bind_thm ("nat_rec_unfold", (wf_pred_nat RS (nat_rec_def RS def_wfrec))); 
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   173
 *---------------------------------------------------------------------------*)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   174
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   175
(** conversion rules **)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   176
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   177
goal thy "nat_rec c h 0 = c";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   178
by (rtac (nat_rec_unfold RS trans) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   179
by (simp_tac (!simpset addsimps [nat_case_0]) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   180
qed "nat_rec_0";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   181
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   182
goal thy "nat_rec c h (Suc n) = h n (nat_rec c h n)";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   183
by (rtac (nat_rec_unfold RS trans) 1);
3236
882e125ed7da New pattern-matching definition of pred_nat
paulson
parents: 3143
diff changeset
   184
by (simp_tac (!simpset addsimps [nat_case_Suc, pred_nat_def, cut_apply]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   185
qed "nat_rec_Suc";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   186
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   187
(*These 2 rules ease the use of primitive recursion.  NOTE USE OF == *)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   188
val [rew] = goal thy
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   189
    "[| !!n. f(n) == nat_rec c h n |] ==> f(0) = c";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   190
by (rewtac rew);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   191
by (rtac nat_rec_0 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   192
qed "def_nat_rec_0";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   193
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   194
val [rew] = goal thy
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   195
    "[| !!n. f(n) == nat_rec c h n |] ==> f(Suc(n)) = h n (f n)";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   196
by (rewtac rew);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   197
by (rtac nat_rec_Suc 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   198
qed "def_nat_rec_Suc";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   199
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   200
fun nat_recs def =
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   201
      [standard (def RS def_nat_rec_0),
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   202
       standard (def RS def_nat_rec_Suc)];
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   203
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   204
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   205
(*** Basic properties of "less than" ***)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   206
3378
11f4884a071a Moved "less_eq" to NatDef from Arith
paulson
parents: 3355
diff changeset
   207
(*Used in TFL/post.sml*)
11f4884a071a Moved "less_eq" to NatDef from Arith
paulson
parents: 3355
diff changeset
   208
goalw thy [less_def] "(m,n) : pred_nat^+ = (m<n)";
11f4884a071a Moved "less_eq" to NatDef from Arith
paulson
parents: 3355
diff changeset
   209
by (rtac refl 1);
11f4884a071a Moved "less_eq" to NatDef from Arith
paulson
parents: 3355
diff changeset
   210
qed "less_eq";
11f4884a071a Moved "less_eq" to NatDef from Arith
paulson
parents: 3355
diff changeset
   211
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   212
(** Introduction properties **)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   213
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   214
val prems = goalw thy [less_def] "[| i<j;  j<k |] ==> i<(k::nat)";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   215
by (rtac (trans_trancl RS transD) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   216
by (resolve_tac prems 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   217
by (resolve_tac prems 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   218
qed "less_trans";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   219
3236
882e125ed7da New pattern-matching definition of pred_nat
paulson
parents: 3143
diff changeset
   220
goalw thy [less_def, pred_nat_def] "n < Suc(n)";
882e125ed7da New pattern-matching definition of pred_nat
paulson
parents: 3143
diff changeset
   221
by (simp_tac (!simpset addsimps [r_into_trancl]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   222
qed "lessI";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   223
AddIffs [lessI];
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   224
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   225
(* i<j ==> i<Suc(j) *)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   226
bind_thm("less_SucI", lessI RSN (2, less_trans));
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   227
Addsimps [less_SucI];
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   228
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   229
goal thy "0 < Suc(n)";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   230
by (nat_ind_tac "n" 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   231
by (rtac lessI 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   232
by (etac less_trans 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   233
by (rtac lessI 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   234
qed "zero_less_Suc";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   235
AddIffs [zero_less_Suc];
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   236
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   237
(** Elimination properties **)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   238
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   239
val prems = goalw thy [less_def] "n<m ==> ~ m<(n::nat)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   240
by (blast_tac (!claset addIs ([wf_pred_nat, wf_trancl RS wf_asym]@prems))1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   241
qed "less_not_sym";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   242
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   243
(* [| n<m; m<n |] ==> R *)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   244
bind_thm ("less_asym", (less_not_sym RS notE));
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   245
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   246
goalw thy [less_def] "~ n<(n::nat)";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   247
by (rtac notI 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   248
by (etac (wf_pred_nat RS wf_trancl RS wf_irrefl) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   249
qed "less_not_refl";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   250
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   251
(* n<n ==> R *)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   252
bind_thm ("less_irrefl", (less_not_refl RS notE));
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   253
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   254
goal thy "!!m. n<m ==> m ~= (n::nat)";
3085
f45074fab9c7 Fixed clasets so that blast_tac would work
paulson
parents: 3040
diff changeset
   255
by (blast_tac (!claset addSEs [less_irrefl]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   256
qed "less_not_refl2";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   257
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   258
3236
882e125ed7da New pattern-matching definition of pred_nat
paulson
parents: 3143
diff changeset
   259
val major::prems = goalw thy [less_def, pred_nat_def]
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   260
    "[| i<k;  k=Suc(i) ==> P;  !!j. [| i<j;  k=Suc(j) |] ==> P \
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   261
\    |] ==> P";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   262
by (rtac (major RS tranclE) 1);
3236
882e125ed7da New pattern-matching definition of pred_nat
paulson
parents: 3143
diff changeset
   263
by (ALLGOALS Full_simp_tac); 
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   264
by (REPEAT_FIRST (bound_hyp_subst_tac ORELSE'
3236
882e125ed7da New pattern-matching definition of pred_nat
paulson
parents: 3143
diff changeset
   265
                  eresolve_tac (prems@[asm_rl, Pair_inject])));
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   266
qed "lessE";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   267
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   268
goal thy "~ n<0";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   269
by (rtac notI 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   270
by (etac lessE 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   271
by (etac Zero_neq_Suc 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   272
by (etac Zero_neq_Suc 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   273
qed "not_less0";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   274
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   275
AddIffs [not_less0];
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   276
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   277
(* n<0 ==> R *)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   278
bind_thm ("less_zeroE", not_less0 RS notE);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   279
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   280
val [major,less,eq] = goal thy
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   281
    "[| m < Suc(n);  m<n ==> P;  m=n ==> P |] ==> P";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   282
by (rtac (major RS lessE) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   283
by (rtac eq 1);
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   284
by (Blast_tac 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   285
by (rtac less 1);
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   286
by (Blast_tac 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   287
qed "less_SucE";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   288
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   289
goal thy "(m < Suc(n)) = (m < n | m = n)";
2935
998cb95fdd43 Yet more fast_tac->blast_tac, and other tidying
paulson
parents: 2922
diff changeset
   290
by (blast_tac (!claset addSEs [less_SucE] addIs [less_trans]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   291
qed "less_Suc_eq";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   292
3484
1e93eb09ebb9 Added the following lemmas tp Divides and a few others to Arith and NatDef:
nipkow
parents: 3457
diff changeset
   293
goal thy "(n<1) = (n=0)";
1e93eb09ebb9 Added the following lemmas tp Divides and a few others to Arith and NatDef:
nipkow
parents: 3457
diff changeset
   294
by (simp_tac (!simpset addsimps [less_Suc_eq]) 1);
1e93eb09ebb9 Added the following lemmas tp Divides and a few others to Arith and NatDef:
nipkow
parents: 3457
diff changeset
   295
qed "less_one";
1e93eb09ebb9 Added the following lemmas tp Divides and a few others to Arith and NatDef:
nipkow
parents: 3457
diff changeset
   296
AddIffs [less_one];
1e93eb09ebb9 Added the following lemmas tp Divides and a few others to Arith and NatDef:
nipkow
parents: 3457
diff changeset
   297
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   298
val prems = goal thy "m<n ==> n ~= 0";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   299
by (res_inst_tac [("n","n")] natE 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   300
by (cut_facts_tac prems 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   301
by (ALLGOALS Asm_full_simp_tac);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   302
qed "gr_implies_not0";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   303
Addsimps [gr_implies_not0];
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   304
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   305
qed_goal "zero_less_eq" thy "0 < n = (n ~= 0)" (fn _ => [
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   306
        rtac iffI 1,
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   307
        etac gr_implies_not0 1,
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   308
        rtac natE 1,
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   309
        contr_tac 1,
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   310
        etac ssubst 1,
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   311
        rtac zero_less_Suc 1]);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   312
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   313
(** Inductive (?) properties **)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   314
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   315
val [prem] = goal thy "Suc(m) < n ==> m<n";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   316
by (rtac (prem RS rev_mp) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   317
by (nat_ind_tac "n" 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   318
by (ALLGOALS (fast_tac (!claset addSIs [lessI RS less_SucI]
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   319
                                addEs  [less_trans, lessE])));
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   320
qed "Suc_lessD";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   321
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   322
val [major,minor] = goal thy 
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   323
    "[| Suc(i)<k;  !!j. [| i<j;  k=Suc(j) |] ==> P \
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   324
\    |] ==> P";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   325
by (rtac (major RS lessE) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   326
by (etac (lessI RS minor) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   327
by (etac (Suc_lessD RS minor) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   328
by (assume_tac 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   329
qed "Suc_lessE";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   330
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   331
goal thy "!!m n. Suc(m) < Suc(n) ==> m<n";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   332
by (blast_tac (!claset addEs [lessE, make_elim Suc_lessD]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   333
qed "Suc_less_SucD";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   334
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   335
goal thy "!!m n. m<n ==> Suc(m) < Suc(n)";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   336
by (etac rev_mp 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   337
by (nat_ind_tac "n" 1);
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   338
by (ALLGOALS (fast_tac (!claset addEs  [less_trans, lessE])));
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   339
qed "Suc_mono";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   340
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   341
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   342
goal thy "(Suc(m) < Suc(n)) = (m<n)";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   343
by (EVERY1 [rtac iffI, etac Suc_less_SucD, etac Suc_mono]);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   344
qed "Suc_less_eq";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   345
Addsimps [Suc_less_eq];
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   346
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   347
goal thy "~(Suc(n) < n)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   348
by (blast_tac (!claset addEs [Suc_lessD RS less_irrefl]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   349
qed "not_Suc_n_less_n";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   350
Addsimps [not_Suc_n_less_n];
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   351
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   352
goal thy "!!i. i<j ==> j<k --> Suc i < k";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   353
by (nat_ind_tac "k" 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   354
by (ALLGOALS (asm_simp_tac (!simpset)));
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   355
by (asm_simp_tac (!simpset addsimps [less_Suc_eq]) 1);
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   356
by (blast_tac (!claset addDs [Suc_lessD]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   357
qed_spec_mp "less_trans_Suc";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   358
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   359
(*"Less than" is a linear ordering*)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   360
goal thy "m<n | m=n | n<(m::nat)";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   361
by (nat_ind_tac "m" 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   362
by (nat_ind_tac "n" 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   363
by (rtac (refl RS disjI1 RS disjI2) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   364
by (rtac (zero_less_Suc RS disjI1) 1);
2935
998cb95fdd43 Yet more fast_tac->blast_tac, and other tidying
paulson
parents: 2922
diff changeset
   365
by (blast_tac (!claset addIs [Suc_mono, less_SucI] addEs [lessE]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   366
qed "less_linear";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   367
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   368
qed_goal "nat_less_cases" thy 
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   369
   "[| (m::nat)<n ==> P n m; m=n ==> P n m; n<m ==> P n m |] ==> P n m"
2935
998cb95fdd43 Yet more fast_tac->blast_tac, and other tidying
paulson
parents: 2922
diff changeset
   370
( fn [major,eqCase,lessCase] =>
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   371
        [
2935
998cb95fdd43 Yet more fast_tac->blast_tac, and other tidying
paulson
parents: 2922
diff changeset
   372
        (rtac (less_linear RS disjE) 1),
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   373
        (etac disjE 2),
2935
998cb95fdd43 Yet more fast_tac->blast_tac, and other tidying
paulson
parents: 2922
diff changeset
   374
        (etac lessCase 1),
998cb95fdd43 Yet more fast_tac->blast_tac, and other tidying
paulson
parents: 2922
diff changeset
   375
        (etac (sym RS eqCase) 1),
998cb95fdd43 Yet more fast_tac->blast_tac, and other tidying
paulson
parents: 2922
diff changeset
   376
        (etac major 1)
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   377
        ]);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   378
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   379
(*Can be used with less_Suc_eq to get n=m | n<m *)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   380
goal thy "(~ m < n) = (n < Suc(m))";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   381
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   382
by (ALLGOALS Asm_simp_tac);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   383
qed "not_less_eq";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   384
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   385
(*Complete induction, aka course-of-values induction*)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   386
val prems = goalw thy [less_def]
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   387
    "[| !!n. [| ! m::nat. m<n --> P(m) |] ==> P(n) |]  ==>  P(n)";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   388
by (wf_ind_tac "n" [wf_pred_nat RS wf_trancl] 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   389
by (eresolve_tac prems 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   390
qed "less_induct";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   391
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   392
qed_goal "nat_induct2" thy 
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   393
"[| P 0; P 1; !!k. P k ==> P (Suc (Suc k)) |] ==> P n" (fn prems => [
3023
01364e2f30ad Ran expandshort
paulson
parents: 2935
diff changeset
   394
        cut_facts_tac prems 1,
01364e2f30ad Ran expandshort
paulson
parents: 2935
diff changeset
   395
        rtac less_induct 1,
01364e2f30ad Ran expandshort
paulson
parents: 2935
diff changeset
   396
        res_inst_tac [("n","n")] natE 1,
01364e2f30ad Ran expandshort
paulson
parents: 2935
diff changeset
   397
         hyp_subst_tac 1,
01364e2f30ad Ran expandshort
paulson
parents: 2935
diff changeset
   398
         atac 1,
01364e2f30ad Ran expandshort
paulson
parents: 2935
diff changeset
   399
        hyp_subst_tac 1,
01364e2f30ad Ran expandshort
paulson
parents: 2935
diff changeset
   400
        res_inst_tac [("n","x")] natE 1,
01364e2f30ad Ran expandshort
paulson
parents: 2935
diff changeset
   401
         hyp_subst_tac 1,
01364e2f30ad Ran expandshort
paulson
parents: 2935
diff changeset
   402
         atac 1,
01364e2f30ad Ran expandshort
paulson
parents: 2935
diff changeset
   403
        hyp_subst_tac 1,
01364e2f30ad Ran expandshort
paulson
parents: 2935
diff changeset
   404
        resolve_tac prems 1,
01364e2f30ad Ran expandshort
paulson
parents: 2935
diff changeset
   405
        dtac spec 1,
01364e2f30ad Ran expandshort
paulson
parents: 2935
diff changeset
   406
        etac mp 1,
01364e2f30ad Ran expandshort
paulson
parents: 2935
diff changeset
   407
        rtac (lessI RS less_trans) 1,
01364e2f30ad Ran expandshort
paulson
parents: 2935
diff changeset
   408
        rtac (lessI RS Suc_mono) 1]);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   409
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   410
(*** Properties of <= ***)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   411
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   412
goalw thy [le_def] "(m <= n) = (m < Suc n)";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   413
by (rtac not_less_eq 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   414
qed "le_eq_less_Suc";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   415
3343
45986997f1fe Renamed lessD to Suc_leI
paulson
parents: 3308
diff changeset
   416
(*  m<=n ==> m < Suc n  *)
45986997f1fe Renamed lessD to Suc_leI
paulson
parents: 3308
diff changeset
   417
bind_thm ("le_imp_less_Suc", le_eq_less_Suc RS iffD1);
45986997f1fe Renamed lessD to Suc_leI
paulson
parents: 3308
diff changeset
   418
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   419
goalw thy [le_def] "0 <= n";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   420
by (rtac not_less0 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   421
qed "le0";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   422
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   423
goalw thy [le_def] "~ Suc n <= n";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   424
by (Simp_tac 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   425
qed "Suc_n_not_le_n";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   426
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   427
goalw thy [le_def] "(i <= 0) = (i = 0)";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   428
by (nat_ind_tac "i" 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   429
by (ALLGOALS Asm_simp_tac);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   430
qed "le_0_eq";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   431
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   432
Addsimps [(*less_Suc_eq, makes simpset non-confluent*) le0, le_0_eq,
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   433
          Suc_n_not_le_n,
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   434
          n_not_Suc_n, Suc_n_not_n,
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   435
          nat_case_0, nat_case_Suc, nat_rec_0, nat_rec_Suc];
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   436
3355
0d955bcf8e0a New theorem le_Suc_eq
paulson
parents: 3343
diff changeset
   437
goal thy "!!m. (m <= Suc(n)) = (m<=n | m = Suc n)";
0d955bcf8e0a New theorem le_Suc_eq
paulson
parents: 3343
diff changeset
   438
by (simp_tac (!simpset addsimps [le_eq_less_Suc]) 1);
0d955bcf8e0a New theorem le_Suc_eq
paulson
parents: 3343
diff changeset
   439
by (blast_tac (!claset addSEs [less_SucE] addIs [less_SucI]) 1);
0d955bcf8e0a New theorem le_Suc_eq
paulson
parents: 3343
diff changeset
   440
qed "le_Suc_eq";
0d955bcf8e0a New theorem le_Suc_eq
paulson
parents: 3343
diff changeset
   441
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   442
(*
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   443
goal thy "(Suc m < n | Suc m = n) = (m < n)";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   444
by (stac (less_Suc_eq RS sym) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   445
by (rtac Suc_less_eq 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   446
qed "Suc_le_eq";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   447
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   448
this could make the simpset (with less_Suc_eq added again) more confluent,
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   449
but less_Suc_eq makes additional problems with terms of the form 0 < Suc (...)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   450
*)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   451
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   452
(*Prevents simplification of f and g: much faster*)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   453
qed_goal "nat_case_weak_cong" thy
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   454
  "m=n ==> nat_case a f m = nat_case a f n"
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   455
  (fn [prem] => [rtac (prem RS arg_cong) 1]);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   456
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   457
qed_goal "nat_rec_weak_cong" thy
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   458
  "m=n ==> nat_rec a f m = nat_rec a f n"
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   459
  (fn [prem] => [rtac (prem RS arg_cong) 1]);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   460
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   461
qed_goal "expand_nat_case" thy
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   462
  "P(nat_case z s n) = ((n=0 --> P z) & (!m. n = Suc m --> P(s m)))"
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   463
  (fn _ => [nat_ind_tac "n" 1, ALLGOALS Asm_simp_tac]);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   464
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   465
val prems = goalw thy [le_def] "~n<m ==> m<=(n::nat)";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   466
by (resolve_tac prems 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   467
qed "leI";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   468
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   469
val prems = goalw thy [le_def] "m<=n ==> ~ n < (m::nat)";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   470
by (resolve_tac prems 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   471
qed "leD";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   472
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   473
val leE = make_elim leD;
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   474
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   475
goal thy "(~n<m) = (m<=(n::nat))";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   476
by (blast_tac (!claset addIs [leI] addEs [leE]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   477
qed "not_less_iff_le";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   478
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   479
goalw thy [le_def] "!!m. ~ m <= n ==> n<(m::nat)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   480
by (Blast_tac 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   481
qed "not_leE";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   482
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   483
goalw thy [le_def] "!!m. m < n ==> Suc(m) <= n";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   484
by (simp_tac (!simpset addsimps [less_Suc_eq]) 1);
3085
f45074fab9c7 Fixed clasets so that blast_tac would work
paulson
parents: 3040
diff changeset
   485
by (blast_tac (!claset addSEs [less_irrefl,less_asym]) 1);
3343
45986997f1fe Renamed lessD to Suc_leI
paulson
parents: 3308
diff changeset
   486
qed "Suc_leI";  (*formerly called lessD*)
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   487
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   488
goalw thy [le_def] "!!m. Suc(m) <= n ==> m <= n";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   489
by (asm_full_simp_tac (!simpset addsimps [less_Suc_eq]) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   490
qed "Suc_leD";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   491
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   492
(* stronger version of Suc_leD *)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   493
goalw thy [le_def] 
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   494
        "!!m. Suc m <= n ==> m < n";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   495
by (asm_full_simp_tac (!simpset addsimps [less_Suc_eq]) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   496
by (cut_facts_tac [less_linear] 1);
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   497
by (Blast_tac 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   498
qed "Suc_le_lessD";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   499
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   500
goal thy "(Suc m <= n) = (m < n)";
3343
45986997f1fe Renamed lessD to Suc_leI
paulson
parents: 3308
diff changeset
   501
by (blast_tac (!claset addIs [Suc_leI, Suc_le_lessD]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   502
qed "Suc_le_eq";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   503
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   504
goalw thy [le_def] "!!m. m <= n ==> m <= Suc n";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   505
by (blast_tac (!claset addDs [Suc_lessD]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   506
qed "le_SucI";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   507
Addsimps[le_SucI];
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   508
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   509
bind_thm ("le_Suc", not_Suc_n_less_n RS leI);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   510
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   511
goalw thy [le_def] "!!m. m < n ==> m <= (n::nat)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   512
by (blast_tac (!claset addEs [less_asym]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   513
qed "less_imp_le";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   514
3343
45986997f1fe Renamed lessD to Suc_leI
paulson
parents: 3308
diff changeset
   515
(** Equivalence of m<=n and  m<n | m=n **)
45986997f1fe Renamed lessD to Suc_leI
paulson
parents: 3308
diff changeset
   516
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   517
goalw thy [le_def] "!!m. m <= n ==> m < n | m=(n::nat)";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   518
by (cut_facts_tac [less_linear] 1);
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   519
by (blast_tac (!claset addEs [less_irrefl,less_asym]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   520
qed "le_imp_less_or_eq";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   521
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   522
goalw thy [le_def] "!!m. m<n | m=n ==> m <=(n::nat)";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   523
by (cut_facts_tac [less_linear] 1);
3085
f45074fab9c7 Fixed clasets so that blast_tac would work
paulson
parents: 3040
diff changeset
   524
by (blast_tac (!claset addSEs [less_irrefl] addEs [less_asym]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   525
qed "less_or_eq_imp_le";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   526
3343
45986997f1fe Renamed lessD to Suc_leI
paulson
parents: 3308
diff changeset
   527
goal thy "(m <= (n::nat)) = (m < n | m=n)";
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   528
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
   529
qed "le_eq_less_or_eq";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   530
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   531
goal thy "n <= (n::nat)";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   532
by (simp_tac (!simpset addsimps [le_eq_less_or_eq]) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   533
qed "le_refl";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   534
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   535
val prems = goal thy "!!i. [| i <= j; j < k |] ==> i < (k::nat)";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   536
by (dtac le_imp_less_or_eq 1);
2935
998cb95fdd43 Yet more fast_tac->blast_tac, and other tidying
paulson
parents: 2922
diff changeset
   537
by (blast_tac (!claset addIs [less_trans]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   538
qed "le_less_trans";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   539
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   540
goal thy "!!i. [| i < j; j <= k |] ==> i < (k::nat)";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   541
by (dtac le_imp_less_or_eq 1);
2935
998cb95fdd43 Yet more fast_tac->blast_tac, and other tidying
paulson
parents: 2922
diff changeset
   542
by (blast_tac (!claset addIs [less_trans]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   543
qed "less_le_trans";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   544
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   545
goal thy "!!i. [| i <= j; j <= k |] ==> i <= (k::nat)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   546
by (EVERY1[dtac le_imp_less_or_eq, 
3023
01364e2f30ad Ran expandshort
paulson
parents: 2935
diff changeset
   547
           dtac le_imp_less_or_eq,
01364e2f30ad Ran expandshort
paulson
parents: 2935
diff changeset
   548
           rtac less_or_eq_imp_le, 
01364e2f30ad Ran expandshort
paulson
parents: 2935
diff changeset
   549
           blast_tac (!claset addIs [less_trans])]);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   550
qed "le_trans";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   551
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   552
goal thy "!!m. [| m <= n; n <= m |] ==> m = (n::nat)";
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   553
by (EVERY1[dtac le_imp_less_or_eq, 
3023
01364e2f30ad Ran expandshort
paulson
parents: 2935
diff changeset
   554
           dtac le_imp_less_or_eq,
01364e2f30ad Ran expandshort
paulson
parents: 2935
diff changeset
   555
           blast_tac (!claset addEs [less_irrefl,less_asym])]);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   556
qed "le_anti_sym";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   557
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   558
goal thy "(Suc(n) <= Suc(m)) = (n <= m)";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   559
by (simp_tac (!simpset addsimps [le_eq_less_or_eq]) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   560
qed "Suc_le_mono";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   561
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   562
AddIffs [Suc_le_mono];
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   563
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   564
(* Axiom 'order_le_less' of class 'order': *)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   565
goal thy "(m::nat) < n = (m <= n & m ~= n)";
3023
01364e2f30ad Ran expandshort
paulson
parents: 2935
diff changeset
   566
by (rtac iffI 1);
01364e2f30ad Ran expandshort
paulson
parents: 2935
diff changeset
   567
 by (rtac conjI 1);
01364e2f30ad Ran expandshort
paulson
parents: 2935
diff changeset
   568
  by (etac less_imp_le 1);
01364e2f30ad Ran expandshort
paulson
parents: 2935
diff changeset
   569
 by (etac (less_not_refl2 RS not_sym) 1);
01364e2f30ad Ran expandshort
paulson
parents: 2935
diff changeset
   570
by (blast_tac (!claset addSDs [le_imp_less_or_eq]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   571
qed "nat_less_le";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   572
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   573
(** LEAST -- the least number operator **)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   574
3143
d60e49b86c6a Modified def of Least, which, as Markus correctly complained, looked like
nipkow
parents: 3085
diff changeset
   575
goal thy "(! m::nat. P m --> n <= m) = (! m. m < n --> ~ P m)";
3457
a8ab7c64817c Ran expandshort
paulson
parents: 3378
diff changeset
   576
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
   577
val lemma = result();
d60e49b86c6a Modified def of Least, which, as Markus correctly complained, looked like
nipkow
parents: 3085
diff changeset
   578
d60e49b86c6a Modified def of Least, which, as Markus correctly complained, looked like
nipkow
parents: 3085
diff changeset
   579
(* This is an old def of Least for nat, which is derived for compatibility *)
d60e49b86c6a Modified def of Least, which, as Markus correctly complained, looked like
nipkow
parents: 3085
diff changeset
   580
goalw thy [Least_def]
d60e49b86c6a Modified def of Least, which, as Markus correctly complained, looked like
nipkow
parents: 3085
diff changeset
   581
  "(LEAST n::nat. P n) == (@n. P(n) & (ALL m. m < n --> ~P(m)))";
3457
a8ab7c64817c Ran expandshort
paulson
parents: 3378
diff changeset
   582
by (simp_tac (!simpset addsimps [lemma]) 1);
a8ab7c64817c Ran expandshort
paulson
parents: 3378
diff changeset
   583
by (rtac eq_reflection 1);
a8ab7c64817c Ran expandshort
paulson
parents: 3378
diff changeset
   584
by (rtac refl 1);
3143
d60e49b86c6a Modified def of Least, which, as Markus correctly complained, looked like
nipkow
parents: 3085
diff changeset
   585
qed "Least_nat_def";
d60e49b86c6a Modified def of Least, which, as Markus correctly complained, looked like
nipkow
parents: 3085
diff changeset
   586
d60e49b86c6a Modified def of Least, which, as Markus correctly complained, looked like
nipkow
parents: 3085
diff changeset
   587
val [prem1,prem2] = goalw thy [Least_nat_def]
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   588
    "[| P(k::nat);  !!x. x<k ==> ~P(x) |] ==> (LEAST x.P(x)) = k";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   589
by (rtac select_equality 1);
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   590
by (blast_tac (!claset addSIs [prem1,prem2]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   591
by (cut_facts_tac [less_linear] 1);
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   592
by (blast_tac (!claset addSIs [prem1] addSDs [prem2]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   593
qed "Least_equality";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   594
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   595
val [prem] = goal thy "P(k::nat) ==> P(LEAST x.P(x))";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   596
by (rtac (prem RS rev_mp) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   597
by (res_inst_tac [("n","k")] less_induct 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   598
by (rtac impI 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   599
by (rtac classical 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   600
by (res_inst_tac [("s","n")] (Least_equality RS ssubst) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   601
by (assume_tac 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   602
by (assume_tac 2);
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   603
by (Blast_tac 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   604
qed "LeastI";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   605
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   606
(*Proof is almost identical to the one above!*)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   607
val [prem] = goal thy "P(k::nat) ==> (LEAST x.P(x)) <= k";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   608
by (rtac (prem RS rev_mp) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   609
by (res_inst_tac [("n","k")] less_induct 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   610
by (rtac impI 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   611
by (rtac classical 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   612
by (res_inst_tac [("s","n")] (Least_equality RS ssubst) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   613
by (assume_tac 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   614
by (rtac le_refl 2);
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   615
by (blast_tac (!claset addIs [less_imp_le,le_trans]) 1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   616
qed "Least_le";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   617
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   618
val [prem] = goal thy "k < (LEAST x.P(x)) ==> ~P(k::nat)";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   619
by (rtac notI 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   620
by (etac (rewrite_rule [le_def] Least_le RS notE) 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   621
by (rtac prem 1);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   622
qed "not_less_Least";
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   623
3143
d60e49b86c6a Modified def of Least, which, as Markus correctly complained, looked like
nipkow
parents: 3085
diff changeset
   624
qed_goalw "Least_Suc" thy [Least_nat_def]
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   625
 "!!P. [| ? n. P(Suc n); ~ P 0 |] ==> (LEAST n. P n) = Suc (LEAST m. P(Suc m))"
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   626
 (fn _ => [
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   627
        rtac select_equality 1,
3143
d60e49b86c6a Modified def of Least, which, as Markus correctly complained, looked like
nipkow
parents: 3085
diff changeset
   628
        fold_goals_tac [Least_nat_def],
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   629
        safe_tac (!claset addSEs [LeastI]),
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   630
        rename_tac "j" 1,
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   631
        res_inst_tac [("n","j")] natE 1,
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   632
        Blast_tac 1,
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   633
        blast_tac (!claset addDs [Suc_less_SucD, not_less_Least]) 1,
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   634
        rename_tac "k n" 1,
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   635
        res_inst_tac [("n","k")] natE 1,
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   636
        Blast_tac 1,
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   637
        hyp_subst_tac 1,
3143
d60e49b86c6a Modified def of Least, which, as Markus correctly complained, looked like
nipkow
parents: 3085
diff changeset
   638
        rewtac Least_nat_def,
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   639
        rtac (select_equality RS arg_cong RS sym) 1,
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   640
        safe_tac (!claset),
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   641
        dtac Suc_mono 1,
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   642
        Blast_tac 1,
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   643
        cut_facts_tac [less_linear] 1,
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   644
        safe_tac (!claset),
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   645
        atac 2,
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   646
        Blast_tac 2,
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   647
        dtac Suc_mono 1,
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   648
        Blast_tac 1]);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   649
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   650
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   651
(*** Instantiation of transitivity prover ***)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   652
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   653
structure Less_Arith =
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   654
struct
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   655
val nat_leI = leI;
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   656
val nat_leD = leD;
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   657
val lessI = lessI;
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   658
val zero_less_Suc = zero_less_Suc;
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   659
val less_reflE = less_irrefl;
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   660
val less_zeroE = less_zeroE;
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   661
val less_incr = Suc_mono;
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   662
val less_decr = Suc_less_SucD;
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   663
val less_incr_rhs = Suc_mono RS Suc_lessD;
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   664
val less_decr_lhs = Suc_lessD;
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   665
val less_trans_Suc = less_trans_Suc;
3343
45986997f1fe Renamed lessD to Suc_leI
paulson
parents: 3308
diff changeset
   666
val leI = Suc_leI RS (Suc_le_mono RS iffD1);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   667
val not_lessI = leI RS leD
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   668
val not_leI = prove_goal thy "!!m::nat. n < m ==> ~ m <= n"
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   669
  (fn _ => [etac swap2 1, etac leD 1]);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   670
val eqI = prove_goal thy "!!m. [| m < Suc n; n < Suc m |] ==> m=n"
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   671
  (fn _ => [etac less_SucE 1,
2935
998cb95fdd43 Yet more fast_tac->blast_tac, and other tidying
paulson
parents: 2922
diff changeset
   672
            blast_tac (!claset addSDs [Suc_less_SucD] addSEs [less_irrefl]
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2718
diff changeset
   673
                              addDs [less_trans_Suc]) 1,
2935
998cb95fdd43 Yet more fast_tac->blast_tac, and other tidying
paulson
parents: 2922
diff changeset
   674
            assume_tac 1]);
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   675
val leD = le_eq_less_Suc RS iffD1;
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   676
val not_lessD = nat_leI RS leD;
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   677
val not_leD = not_leE
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   678
val eqD1 = prove_goal thy  "!!n. m = n ==> m < Suc n"
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   679
 (fn _ => [etac subst 1, rtac lessI 1]);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   680
val eqD2 = sym RS eqD1;
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   681
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   682
fun is_zero(t) =  t = Const("0",Type("nat",[]));
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   683
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   684
fun nnb T = T = Type("fun",[Type("nat",[]),
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   685
                            Type("fun",[Type("nat",[]),
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   686
                                        Type("bool",[])])])
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   687
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   688
fun decomp_Suc(Const("Suc",_)$t) = let val (a,i) = decomp_Suc t in (a,i+1) end
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   689
  | decomp_Suc t = (t,0);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   690
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   691
fun decomp2(rel,T,lhs,rhs) =
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   692
  if not(nnb T) then None else
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   693
  let val (x,i) = decomp_Suc lhs
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   694
      val (y,j) = decomp_Suc rhs
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   695
  in case rel of
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   696
       "op <"  => Some(x,i,"<",y,j)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   697
     | "op <=" => Some(x,i,"<=",y,j)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   698
     | "op ="  => Some(x,i,"=",y,j)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   699
     | _       => None
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   700
  end;
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   701
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   702
fun negate(Some(x,i,rel,y,j)) = Some(x,i,"~"^rel,y,j)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   703
  | negate None = None;
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   704
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   705
fun decomp(_$(Const(rel,T)$lhs$rhs)) = decomp2(rel,T,lhs,rhs)
2718
460fd0f8d478 Renamed constant "not" to "Not"
paulson
parents: 2680
diff changeset
   706
  | decomp(_$(Const("Not",_)$(Const(rel,T)$lhs$rhs))) =
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   707
      negate(decomp2(rel,T,lhs,rhs))
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   708
  | decomp _ = None
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   709
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   710
end;
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   711
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   712
structure Trans_Tac = Trans_Tac_Fun(Less_Arith);
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   713
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   714
open Trans_Tac;
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   715
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   716
(*** eliminates ~= in premises, which trans_tac cannot deal with ***)
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   717
qed_goal "nat_neqE" thy
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   718
  "[| (m::nat) ~= n; m < n ==> P; n < m ==> P |] ==> P"
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   719
  (fn major::prems => [cut_facts_tac [less_linear] 1,
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents:
diff changeset
   720
                       REPEAT(eresolve_tac ([disjE,major RS notE]@prems) 1)]);
2680
20fa49e610ca function nat_add_primrec added to allow primrec definitions over nat
pusch
parents: 2608
diff changeset
   721
20fa49e610ca function nat_add_primrec added to allow primrec definitions over nat
pusch
parents: 2608
diff changeset
   722
20fa49e610ca function nat_add_primrec added to allow primrec definitions over nat
pusch
parents: 2608
diff changeset
   723
20fa49e610ca function nat_add_primrec added to allow primrec definitions over nat
pusch
parents: 2608
diff changeset
   724
(* add function nat_add_primrec *) 
3308
da002cef7090 Added overloaded function `size' for all datatypes.
nipkow
parents: 3292
diff changeset
   725
val (_, nat_add_primrec, _) = Datatype.add_datatype
da002cef7090 Added overloaded function `size' for all datatypes.
nipkow
parents: 3292
diff changeset
   726
([], "nat", [("0", [], Mixfix ("0", [], max_pri)), ("Suc", [dtTyp ([],
da002cef7090 Added overloaded function `size' for all datatypes.
nipkow
parents: 3292
diff changeset
   727
"nat")], NoSyn)]) (add_thyname "Arith" HOL.thy);
da002cef7090 Added overloaded function `size' for all datatypes.
nipkow
parents: 3292
diff changeset
   728
(* pretend Arith is part of the basic theory to fool package *)
2680
20fa49e610ca function nat_add_primrec added to allow primrec definitions over nat
pusch
parents: 2608
diff changeset
   729