src/HOL/Nat.ML
author berghofe
Fri, 21 Jun 1996 13:51:09 +0200
changeset 1823 e1458e1a9f80
parent 1817 3994d37b16ae
child 1824 44254696843a
permissions -rw-r--r--
Replaced occurrence of fast_tac by Fast_tac .
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1327
diff changeset
     1
(*  Title:      HOL/nat
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1327
diff changeset
     3
    Author:     Tobias Nipkow, Cambridge University Computer Laboratory
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     4
    Copyright   1991  University of Cambridge
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     5
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     6
For nat.thy.  Type nat is defined as a set (Nat) over the type ind.
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     7
*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     8
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     9
open Nat;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    10
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    11
goal Nat.thy "mono(%X. {Zero_Rep} Un (Suc_Rep``X))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    12
by (REPEAT (ares_tac [monoI, subset_refl, image_mono, Un_mono] 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    13
qed "Nat_fun_mono";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    14
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    15
val Nat_unfold = Nat_fun_mono RS (Nat_def RS def_lfp_Tarski);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    16
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    17
(* Zero is a natural number -- this also justifies the type definition*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    18
goal Nat.thy "Zero_Rep: Nat";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    19
by (rtac (Nat_unfold RS ssubst) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    20
by (rtac (singletonI RS UnI1) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    21
qed "Zero_RepI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    22
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    23
val prems = goal Nat.thy "i: Nat ==> Suc_Rep(i) : Nat";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    24
by (rtac (Nat_unfold RS ssubst) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    25
by (rtac (imageI RS UnI2) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    26
by (resolve_tac prems 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    27
qed "Suc_RepI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    28
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    29
(*** Induction ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    30
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    31
val major::prems = goal Nat.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    32
    "[| i: Nat;  P(Zero_Rep);   \
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    33
\       !!j. [| j: Nat; P(j) |] ==> P(Suc_Rep(j)) |]  ==> P(i)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    34
by (rtac ([Nat_def, Nat_fun_mono, major] MRS def_induct) 1);
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
    35
by (fast_tac (!claset addIs prems) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    36
qed "Nat_induct";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    37
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    38
val prems = goalw Nat.thy [Zero_def,Suc_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    39
    "[| P(0);   \
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    40
\       !!k. P(k) ==> P(Suc(k)) |]  ==> P(n)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    41
by (rtac (Rep_Nat_inverse RS subst) 1);   (*types force good instantiation*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    42
by (rtac (Rep_Nat RS Nat_induct) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    43
by (REPEAT (ares_tac prems 1
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    44
     ORELSE eresolve_tac [Abs_Nat_inverse RS subst] 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    45
qed "nat_induct";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    46
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    47
(*Perform induction on n. *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    48
fun nat_ind_tac a i = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    49
    EVERY [res_inst_tac [("n",a)] nat_induct i,
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1327
diff changeset
    50
           rename_last_tac a ["1"] (i+1)];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    51
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    52
(*A special form of induction for reasoning about m<n and m-n*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    53
val prems = goal Nat.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    54
    "[| !!x. P x 0;  \
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    55
\       !!y. P 0 (Suc y);  \
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    56
\       !!x y. [| P x y |] ==> P (Suc x) (Suc y)  \
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    57
\    |] ==> P m n";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    58
by (res_inst_tac [("x","m")] spec 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    59
by (nat_ind_tac "n" 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    60
by (rtac allI 2);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    61
by (nat_ind_tac "x" 2);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    62
by (REPEAT (ares_tac (prems@[allI]) 1 ORELSE etac spec 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    63
qed "diff_induct";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    64
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    65
(*Case analysis on the natural numbers*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    66
val prems = goal Nat.thy 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    67
    "[| n=0 ==> P;  !!x. n = Suc(x) ==> P |] ==> P";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    68
by (subgoal_tac "n=0 | (EX x. n = Suc(x))" 1);
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
    69
by (fast_tac (!claset addSEs prems) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    70
by (nat_ind_tac "n" 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    71
by (rtac (refl RS disjI1) 1);
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
    72
by (Fast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    73
qed "natE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    74
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    75
(*** Isomorphisms: Abs_Nat and Rep_Nat ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    76
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    77
(*We can't take these properties as axioms, or take Abs_Nat==Inv(Rep_Nat),
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    78
  since we assume the isomorphism equations will one day be given by Isabelle*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    79
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    80
goal Nat.thy "inj(Rep_Nat)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    81
by (rtac inj_inverseI 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    82
by (rtac Rep_Nat_inverse 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    83
qed "inj_Rep_Nat";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    84
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    85
goal Nat.thy "inj_onto Abs_Nat Nat";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    86
by (rtac inj_onto_inverseI 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    87
by (etac Abs_Nat_inverse 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    88
qed "inj_onto_Abs_Nat";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    89
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    90
(*** Distinctness of constructors ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    91
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    92
goalw Nat.thy [Zero_def,Suc_def] "Suc(m) ~= 0";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    93
by (rtac (inj_onto_Abs_Nat RS inj_onto_contraD) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    94
by (rtac Suc_Rep_not_Zero_Rep 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    95
by (REPEAT (resolve_tac [Rep_Nat, Suc_RepI, Zero_RepI] 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    96
qed "Suc_not_Zero";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    97
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    98
bind_thm ("Zero_not_Suc", (Suc_not_Zero RS not_sym));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    99
1301
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
   100
Addsimps [Suc_not_Zero,Zero_not_Suc];
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
   101
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   102
bind_thm ("Suc_neq_Zero", (Suc_not_Zero RS notE));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   103
val Zero_neq_Suc = sym RS Suc_neq_Zero;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   104
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   105
(** Injectiveness of Suc **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   106
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   107
goalw Nat.thy [Suc_def] "inj(Suc)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   108
by (rtac injI 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   109
by (dtac (inj_onto_Abs_Nat RS inj_ontoD) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   110
by (REPEAT (resolve_tac [Rep_Nat, Suc_RepI] 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   111
by (dtac (inj_Suc_Rep RS injD) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   112
by (etac (inj_Rep_Nat RS injD) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   113
qed "inj_Suc";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   114
1264
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 1024
diff changeset
   115
val Suc_inject = inj_Suc RS injD;
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   116
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   117
goal Nat.thy "(Suc(m)=Suc(n)) = (m=n)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   118
by (EVERY1 [rtac iffI, etac Suc_inject, etac arg_cong]); 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   119
qed "Suc_Suc_eq";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   120
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   121
goal Nat.thy "n ~= Suc(n)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   122
by (nat_ind_tac "n" 1);
1301
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
   123
by (ALLGOALS(asm_simp_tac (!simpset addsimps [Suc_Suc_eq])));
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   124
qed "n_not_Suc_n";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   125
1618
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1552
diff changeset
   126
bind_thm ("Suc_n_not_n", n_not_Suc_n RS not_sym);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   127
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   128
(*** nat_case -- the selection operator for nat ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   129
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   130
goalw Nat.thy [nat_case_def] "nat_case a f 0 = a";
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   131
by (fast_tac (!claset addIs [select_equality] addEs [Zero_neq_Suc]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   132
qed "nat_case_0";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   133
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   134
goalw Nat.thy [nat_case_def] "nat_case a f (Suc k) = f(k)";
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   135
by (fast_tac (!claset addIs [select_equality] 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1327
diff changeset
   136
                       addEs [make_elim Suc_inject, Suc_neq_Zero]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   137
qed "nat_case_Suc";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   138
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   139
(** Introduction rules for 'pred_nat' **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   140
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 962
diff changeset
   141
goalw Nat.thy [pred_nat_def] "(n, Suc(n)) : pred_nat";
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   142
by (Fast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   143
qed "pred_natI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   144
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   145
val major::prems = goalw Nat.thy [pred_nat_def]
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 962
diff changeset
   146
    "[| p : pred_nat;  !!x n. [| p = (n, Suc(n)) |] ==> R \
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   147
\    |] ==> R";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   148
by (rtac (major RS CollectE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   149
by (REPEAT (eresolve_tac ([asm_rl,exE]@prems) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   150
qed "pred_natE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   151
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   152
goalw Nat.thy [wf_def] "wf(pred_nat)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   153
by (strip_tac 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   154
by (nat_ind_tac "x" 1);
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   155
by (fast_tac (!claset addSEs [mp, pred_natE, Pair_inject, 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1327
diff changeset
   156
                             make_elim Suc_inject]) 2);
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   157
by (fast_tac (!claset addSEs [mp, pred_natE, Pair_inject, Zero_neq_Suc]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   158
qed "wf_pred_nat";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   159
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   160
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   161
(*** nat_rec -- by wf recursion on pred_nat ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   162
1475
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1465
diff changeset
   163
(* The unrolling rule for nat_rec *)
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1465
diff changeset
   164
goal Nat.thy
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1465
diff changeset
   165
   "(%n. nat_rec n c d) = wfrec pred_nat (%f. nat_case ?c (%m. ?d m (f m)))";
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1465
diff changeset
   166
  by (simp_tac (HOL_ss addsimps [nat_rec_def]) 1);
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1465
diff changeset
   167
bind_thm("nat_rec_unfold", wf_pred_nat RS 
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1465
diff changeset
   168
                            ((result() RS eq_reflection) RS def_wfrec));
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1465
diff changeset
   169
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1465
diff changeset
   170
(*---------------------------------------------------------------------------
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1465
diff changeset
   171
 * Old:
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1465
diff changeset
   172
 * bind_thm ("nat_rec_unfold", (wf_pred_nat RS (nat_rec_def RS def_wfrec))); 
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1465
diff changeset
   173
 *---------------------------------------------------------------------------*)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   174
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   175
(** conversion rules **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   176
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   177
goal Nat.thy "nat_rec 0 c h = c";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   178
by (rtac (nat_rec_unfold RS trans) 1);
1264
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 1024
diff changeset
   179
by (simp_tac (!simpset addsimps [nat_case_0]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   180
qed "nat_rec_0";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   181
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   182
goal Nat.thy "nat_rec (Suc n) c h = h n (nat_rec n c h)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   183
by (rtac (nat_rec_unfold RS trans) 1);
1264
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 1024
diff changeset
   184
by (simp_tac (!simpset addsimps [nat_case_Suc, pred_natI, cut_apply]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   185
qed "nat_rec_Suc";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   186
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   187
(*These 2 rules ease the use of primitive recursion.  NOTE USE OF == *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   188
val [rew] = goal Nat.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   189
    "[| !!n. f(n) == nat_rec n c h |] ==> f(0) = c";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   190
by (rewtac rew);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   191
by (rtac nat_rec_0 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   192
qed "def_nat_rec_0";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   193
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   194
val [rew] = goal Nat.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   195
    "[| !!n. f(n) == nat_rec n c h |] ==> f(Suc(n)) = h n (f n)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   196
by (rewtac rew);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   197
by (rtac nat_rec_Suc 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   198
qed "def_nat_rec_Suc";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   199
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   200
fun nat_recs def =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   201
      [standard (def RS def_nat_rec_0),
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   202
       standard (def RS def_nat_rec_Suc)];
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   203
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   204
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   205
(*** Basic properties of "less than" ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   206
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   207
(** Introduction properties **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   208
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   209
val prems = goalw Nat.thy [less_def] "[| i<j;  j<k |] ==> i<(k::nat)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   210
by (rtac (trans_trancl RS transD) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   211
by (resolve_tac prems 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   212
by (resolve_tac prems 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   213
qed "less_trans";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   214
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   215
goalw Nat.thy [less_def] "n < Suc(n)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   216
by (rtac (pred_natI RS r_into_trancl) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   217
qed "lessI";
1301
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
   218
Addsimps [lessI];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   219
1618
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1552
diff changeset
   220
(* i<j ==> i<Suc(j) *)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   221
val less_SucI = lessI RSN (2, less_trans);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   222
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   223
goal Nat.thy "0 < Suc(n)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   224
by (nat_ind_tac "n" 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   225
by (rtac lessI 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   226
by (etac less_trans 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   227
by (rtac lessI 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   228
qed "zero_less_Suc";
1301
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
   229
Addsimps [zero_less_Suc];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   230
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   231
(** Elimination properties **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   232
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   233
val prems = goalw Nat.thy [less_def] "n<m ==> ~ m<(n::nat)";
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   234
by (fast_tac (!claset addIs ([wf_pred_nat, wf_trancl RS wf_asym]@prems))1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   235
qed "less_not_sym";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   236
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 962
diff changeset
   237
(* [| n(m; m(n |] ==> R *)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   238
bind_thm ("less_asym", (less_not_sym RS notE));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   239
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   240
goalw Nat.thy [less_def] "~ n<(n::nat)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   241
by (rtac notI 1);
1618
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1552
diff changeset
   242
by (etac (wf_pred_nat RS wf_trancl RS wf_irrefl) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   243
qed "less_not_refl";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   244
1817
3994d37b16ae Corrected comment
paulson
parents: 1786
diff changeset
   245
(* n<n ==> R *)
1618
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1552
diff changeset
   246
bind_thm ("less_irrefl", (less_not_refl RS notE));
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   247
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   248
goal Nat.thy "!!m. n<m ==> m ~= (n::nat)";
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   249
by (fast_tac (!claset addEs [less_irrefl]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   250
qed "less_not_refl2";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   251
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   252
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   253
val major::prems = goalw Nat.thy [less_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   254
    "[| i<k;  k=Suc(i) ==> P;  !!j. [| i<j;  k=Suc(j) |] ==> P \
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   255
\    |] ==> P";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   256
by (rtac (major RS tranclE) 1);
1024
b86042000035 ROOT.ML: installed new hyp_subst_tac
nipkow
parents: 972
diff changeset
   257
by (REPEAT_FIRST (bound_hyp_subst_tac ORELSE'
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1327
diff changeset
   258
                  eresolve_tac (prems@[pred_natE, Pair_inject])));
1024
b86042000035 ROOT.ML: installed new hyp_subst_tac
nipkow
parents: 972
diff changeset
   259
by (rtac refl 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   260
qed "lessE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   261
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   262
goal Nat.thy "~ n<0";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   263
by (rtac notI 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   264
by (etac lessE 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   265
by (etac Zero_neq_Suc 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   266
by (etac Zero_neq_Suc 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   267
qed "not_less0";
1301
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
   268
Addsimps [not_less0];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   269
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   270
(* n<0 ==> R *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   271
bind_thm ("less_zeroE", (not_less0 RS notE));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   272
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   273
val [major,less,eq] = goal Nat.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   274
    "[| m < Suc(n);  m<n ==> P;  m=n ==> P |] ==> P";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   275
by (rtac (major RS lessE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   276
by (rtac eq 1);
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   277
by (fast_tac (!claset addSDs [Suc_inject]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   278
by (rtac less 1);
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   279
by (fast_tac (!claset addSDs [Suc_inject]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   280
qed "less_SucE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   281
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   282
goal Nat.thy "(m < Suc(n)) = (m < n | m = n)";
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   283
by (fast_tac (!claset addSIs [lessI]
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1327
diff changeset
   284
                     addEs  [less_trans, less_SucE]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   285
qed "less_Suc_eq";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   286
1301
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
   287
val prems = goal Nat.thy "m<n ==> n ~= 0";
1552
6f71b5d46700 Ran expandshort
paulson
parents: 1531
diff changeset
   288
by (res_inst_tac [("n","n")] natE 1);
6f71b5d46700 Ran expandshort
paulson
parents: 1531
diff changeset
   289
by (cut_facts_tac prems 1);
6f71b5d46700 Ran expandshort
paulson
parents: 1531
diff changeset
   290
by (ALLGOALS Asm_full_simp_tac);
1301
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
   291
qed "gr_implies_not0";
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
   292
Addsimps [gr_implies_not0];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   293
1660
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   294
qed_goal "zero_less_eq" Nat.thy "0 < n = (n ~= 0)" (fn _ => [
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   295
	rtac iffI 1,
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   296
	etac gr_implies_not0 1,
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   297
	rtac natE 1,
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   298
	contr_tac 1,
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   299
	etac ssubst 1,
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   300
	rtac zero_less_Suc 1]);
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   301
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   302
(** Inductive (?) properties **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   303
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   304
val [prem] = goal Nat.thy "Suc(m) < n ==> m<n";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   305
by (rtac (prem RS rev_mp) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   306
by (nat_ind_tac "n" 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   307
by (rtac impI 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   308
by (etac less_zeroE 1);
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   309
by (fast_tac (!claset addSIs [lessI RS less_SucI]
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1327
diff changeset
   310
                     addSDs [Suc_inject]
5d7a7e439cec expanded tabs
clasohm
parents: 1327
diff changeset
   311
                     addEs  [less_trans, lessE]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   312
qed "Suc_lessD";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   313
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   314
val [major,minor] = goal Nat.thy 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   315
    "[| Suc(i)<k;  !!j. [| i<j;  k=Suc(j) |] ==> P \
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   316
\    |] ==> P";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   317
by (rtac (major RS lessE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   318
by (etac (lessI RS minor) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   319
by (etac (Suc_lessD RS minor) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   320
by (assume_tac 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   321
qed "Suc_lessE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   322
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   323
val [major] = goal Nat.thy "Suc(m) < Suc(n) ==> m<n";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   324
by (rtac (major RS lessE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   325
by (REPEAT (rtac lessI 1
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   326
     ORELSE eresolve_tac [make_elim Suc_inject, ssubst, Suc_lessD] 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   327
qed "Suc_less_SucD";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   328
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   329
val prems = goal Nat.thy "m<n ==> Suc(m) < Suc(n)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   330
by (subgoal_tac "m<n --> Suc(m) < Suc(n)" 1);
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   331
by (fast_tac (!claset addIs prems) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   332
by (nat_ind_tac "n" 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   333
by (rtac impI 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   334
by (etac less_zeroE 1);
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   335
by (fast_tac (!claset addSIs [lessI]
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1327
diff changeset
   336
                     addSDs [Suc_inject]
5d7a7e439cec expanded tabs
clasohm
parents: 1327
diff changeset
   337
                     addEs  [less_trans, lessE]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   338
qed "Suc_mono";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   339
1672
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1660
diff changeset
   340
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   341
goal Nat.thy "(Suc(m) < Suc(n)) = (m<n)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   342
by (EVERY1 [rtac iffI, etac Suc_less_SucD, etac Suc_mono]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   343
qed "Suc_less_eq";
1301
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
   344
Addsimps [Suc_less_eq];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   345
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   346
goal Nat.thy "~(Suc(n) < n)";
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   347
by (fast_tac (!claset addEs [Suc_lessD RS less_irrefl]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   348
qed "not_Suc_n_less_n";
1301
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
   349
Addsimps [not_Suc_n_less_n];
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
   350
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
   351
goal Nat.thy "!!i. i<j ==> j<k --> Suc i < k";
1552
6f71b5d46700 Ran expandshort
paulson
parents: 1531
diff changeset
   352
by (nat_ind_tac "k" 1);
1660
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   353
by (ALLGOALS (asm_simp_tac (!simpset)));
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   354
by (asm_simp_tac (!simpset addsimps [less_Suc_eq]) 1);
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   355
by (fast_tac (!claset addDs [Suc_lessD]) 1);
1485
240cc98b94a7 Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
nipkow
parents: 1475
diff changeset
   356
qed_spec_mp "less_trans_Suc";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   357
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   358
(*"Less than" is a linear ordering*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   359
goal Nat.thy "m<n | m=n | n<(m::nat)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   360
by (nat_ind_tac "m" 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   361
by (nat_ind_tac "n" 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   362
by (rtac (refl RS disjI1 RS disjI2) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   363
by (rtac (zero_less_Suc RS disjI1) 1);
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   364
by (fast_tac (!claset addIs [lessI, Suc_mono, less_SucI] addEs [lessE]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   365
qed "less_linear";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   366
1660
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   367
qed_goal "nat_less_cases" Nat.thy 
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   368
   "[| (m::nat)<n ==> P n m; m=n ==> P n m; n<m ==> P n m |] ==> P n m"
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   369
( fn prems =>
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   370
        [
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   371
        (res_inst_tac [("m1","n"),("n1","m")] (less_linear RS disjE) 1),
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   372
        (etac disjE 2),
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   373
        (etac (hd (tl (tl prems))) 1),
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   374
        (etac (sym RS hd (tl prems)) 1),
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   375
        (etac (hd prems) 1)
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   376
        ]);
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   377
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   378
(*Can be used with less_Suc_eq to get n=m | n<m *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   379
goal Nat.thy "(~ m < n) = (n < Suc(m))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   380
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
1552
6f71b5d46700 Ran expandshort
paulson
parents: 1531
diff changeset
   381
by (ALLGOALS Asm_simp_tac);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   382
qed "not_less_eq";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   383
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   384
(*Complete induction, aka course-of-values induction*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   385
val prems = goalw Nat.thy [less_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   386
    "[| !!n. [| ! m::nat. m<n --> P(m) |] ==> P(n) |]  ==>  P(n)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   387
by (wf_ind_tac "n" [wf_pred_nat RS wf_trancl] 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   388
by (eresolve_tac prems 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   389
qed "less_induct";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   390
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   391
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   392
(*** Properties of <= ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   393
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   394
goalw Nat.thy [le_def] "0 <= n";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   395
by (rtac not_less0 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   396
qed "le0";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   397
1301
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
   398
goalw Nat.thy [le_def] "~ Suc n <= n";
1552
6f71b5d46700 Ran expandshort
paulson
parents: 1531
diff changeset
   399
by (Simp_tac 1);
1301
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
   400
qed "Suc_n_not_le_n";
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
   401
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
   402
goalw Nat.thy [le_def] "(i <= 0) = (i = 0)";
1552
6f71b5d46700 Ran expandshort
paulson
parents: 1531
diff changeset
   403
by (nat_ind_tac "i" 1);
6f71b5d46700 Ran expandshort
paulson
parents: 1531
diff changeset
   404
by (ALLGOALS Asm_simp_tac);
1777
47c47b7d7aa8 renamed le_0 to le_0_eq, to avoid confusion with le0,
oheimb
parents: 1760
diff changeset
   405
qed "le_0_eq";
1301
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
   406
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
   407
Addsimps [less_not_refl,
1777
47c47b7d7aa8 renamed le_0 to le_0_eq, to avoid confusion with le0,
oheimb
parents: 1760
diff changeset
   408
          (*less_Suc_eq, makes simpset non-confluent*) le0, le_0_eq,
1301
42782316d510 Added various thms and tactics.
nipkow
parents: 1264
diff changeset
   409
          Suc_Suc_eq, Suc_n_not_le_n,
1264
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 1024
diff changeset
   410
          n_not_Suc_n, Suc_n_not_n,
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 1024
diff changeset
   411
          nat_case_0, nat_case_Suc, nat_rec_0, nat_rec_Suc];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   412
1777
47c47b7d7aa8 renamed le_0 to le_0_eq, to avoid confusion with le0,
oheimb
parents: 1760
diff changeset
   413
(*
47c47b7d7aa8 renamed le_0 to le_0_eq, to avoid confusion with le0,
oheimb
parents: 1760
diff changeset
   414
goal Nat.thy "(Suc m < n | Suc m = n) = (m < n)";
47c47b7d7aa8 renamed le_0 to le_0_eq, to avoid confusion with le0,
oheimb
parents: 1760
diff changeset
   415
by(stac (less_Suc_eq RS sym) 1);
47c47b7d7aa8 renamed le_0 to le_0_eq, to avoid confusion with le0,
oheimb
parents: 1760
diff changeset
   416
by(rtac Suc_less_eq 1);
47c47b7d7aa8 renamed le_0 to le_0_eq, to avoid confusion with le0,
oheimb
parents: 1760
diff changeset
   417
qed "Suc_le_eq";
47c47b7d7aa8 renamed le_0 to le_0_eq, to avoid confusion with le0,
oheimb
parents: 1760
diff changeset
   418
47c47b7d7aa8 renamed le_0 to le_0_eq, to avoid confusion with le0,
oheimb
parents: 1760
diff changeset
   419
this could make the simpset (with less_Suc_eq added again) more confluent,
47c47b7d7aa8 renamed le_0 to le_0_eq, to avoid confusion with le0,
oheimb
parents: 1760
diff changeset
   420
but less_Suc_eq makes additional problems with terms of the form 0 < Suc (...)
47c47b7d7aa8 renamed le_0 to le_0_eq, to avoid confusion with le0,
oheimb
parents: 1760
diff changeset
   421
*)
47c47b7d7aa8 renamed le_0 to le_0_eq, to avoid confusion with le0,
oheimb
parents: 1760
diff changeset
   422
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   423
(*Prevents simplification of f and g: much faster*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   424
qed_goal "nat_case_weak_cong" Nat.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   425
  "m=n ==> nat_case a f m = nat_case a f n"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   426
  (fn [prem] => [rtac (prem RS arg_cong) 1]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   427
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   428
qed_goal "nat_rec_weak_cong" Nat.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   429
  "m=n ==> nat_rec m a f = nat_rec n a f"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   430
  (fn [prem] => [rtac (prem RS arg_cong) 1]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   431
1618
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1552
diff changeset
   432
val prems = goalw Nat.thy [le_def] "~n<m ==> m<=(n::nat)";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   433
by (resolve_tac prems 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   434
qed "leI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   435
1618
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1552
diff changeset
   436
val prems = goalw Nat.thy [le_def] "m<=n ==> ~ n < (m::nat)";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   437
by (resolve_tac prems 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   438
qed "leD";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   439
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   440
val leE = make_elim leD;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   441
1618
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1552
diff changeset
   442
goal Nat.thy "(~n<m) = (m<=(n::nat))";
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   443
by (fast_tac (!claset addIs [leI] addEs [leE]) 1);
1618
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1552
diff changeset
   444
qed "not_less_iff_le";
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1552
diff changeset
   445
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   446
goalw Nat.thy [le_def] "!!m. ~ m <= n ==> n<(m::nat)";
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   447
by (Fast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   448
qed "not_leE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   449
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   450
goalw Nat.thy [le_def] "!!m. m < n ==> Suc(m) <= n";
1660
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   451
by (simp_tac (!simpset addsimps [less_Suc_eq]) 1);
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   452
by (fast_tac (!claset addEs [less_irrefl,less_asym]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   453
qed "lessD";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   454
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   455
goalw Nat.thy [le_def] "!!m. Suc(m) <= n ==> m <= n";
1660
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   456
by (asm_full_simp_tac (!simpset addsimps [less_Suc_eq]) 1);
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   457
by (Fast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   458
qed "Suc_leD";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   459
1817
3994d37b16ae Corrected comment
paulson
parents: 1786
diff changeset
   460
(* stronger version of Suc_leD *)
3994d37b16ae Corrected comment
paulson
parents: 1786
diff changeset
   461
goalw Nat.thy [le_def] 
3994d37b16ae Corrected comment
paulson
parents: 1786
diff changeset
   462
        "!!m. Suc m <= n ==> m < n";
3994d37b16ae Corrected comment
paulson
parents: 1786
diff changeset
   463
by (asm_full_simp_tac (!simpset addsimps [less_Suc_eq]) 1);
3994d37b16ae Corrected comment
paulson
parents: 1786
diff changeset
   464
by (cut_facts_tac [less_linear] 1);
1823
e1458e1a9f80 Replaced occurrence of fast_tac by Fast_tac .
berghofe
parents: 1817
diff changeset
   465
by (Fast_tac 1);
1817
3994d37b16ae Corrected comment
paulson
parents: 1786
diff changeset
   466
qed "Suc_le_lessD";
3994d37b16ae Corrected comment
paulson
parents: 1786
diff changeset
   467
3994d37b16ae Corrected comment
paulson
parents: 1786
diff changeset
   468
goal Nat.thy "(Suc m <= n) = (m < n)";
3994d37b16ae Corrected comment
paulson
parents: 1786
diff changeset
   469
by (fast_tac (!claset addIs [lessD, Suc_le_lessD]) 1);
3994d37b16ae Corrected comment
paulson
parents: 1786
diff changeset
   470
qed "Suc_le_eq";
3994d37b16ae Corrected comment
paulson
parents: 1786
diff changeset
   471
1327
6c29cfab679c added new arithmetic lemmas and the functions take and drop.
nipkow
parents: 1301
diff changeset
   472
goalw Nat.thy [le_def] "!!m. m <= n ==> m <= Suc n";
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   473
by (fast_tac (!claset addDs [Suc_lessD]) 1);
1327
6c29cfab679c added new arithmetic lemmas and the functions take and drop.
nipkow
parents: 1301
diff changeset
   474
qed "le_SucI";
6c29cfab679c added new arithmetic lemmas and the functions take and drop.
nipkow
parents: 1301
diff changeset
   475
Addsimps[le_SucI];
6c29cfab679c added new arithmetic lemmas and the functions take and drop.
nipkow
parents: 1301
diff changeset
   476
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   477
goalw Nat.thy [le_def] "!!m. m < n ==> m <= (n::nat)";
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   478
by (fast_tac (!claset addEs [less_asym]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   479
qed "less_imp_le";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   480
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   481
goalw Nat.thy [le_def] "!!m. m <= n ==> m < n | m=(n::nat)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   482
by (cut_facts_tac [less_linear] 1);
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   483
by (fast_tac (!claset addEs [less_irrefl,less_asym]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   484
qed "le_imp_less_or_eq";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   485
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   486
goalw Nat.thy [le_def] "!!m. m<n | m=n ==> m <=(n::nat)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   487
by (cut_facts_tac [less_linear] 1);
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   488
by (fast_tac (!claset addEs [less_irrefl,less_asym]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   489
by (flexflex_tac);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   490
qed "less_or_eq_imp_le";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   491
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   492
goal Nat.thy "(x <= (y::nat)) = (x < y | x=y)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   493
by (REPEAT(ares_tac [iffI,less_or_eq_imp_le,le_imp_less_or_eq] 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   494
qed "le_eq_less_or_eq";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   495
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   496
goal Nat.thy "n <= (n::nat)";
1552
6f71b5d46700 Ran expandshort
paulson
parents: 1531
diff changeset
   497
by (simp_tac (!simpset addsimps [le_eq_less_or_eq]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   498
qed "le_refl";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   499
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   500
val prems = goal Nat.thy "!!i. [| i <= j; j < k |] ==> i < (k::nat)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   501
by (dtac le_imp_less_or_eq 1);
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   502
by (fast_tac (!claset addIs [less_trans]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   503
qed "le_less_trans";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   504
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   505
goal Nat.thy "!!i. [| i < j; j <= k |] ==> i < (k::nat)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   506
by (dtac le_imp_less_or_eq 1);
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   507
by (fast_tac (!claset addIs [less_trans]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   508
qed "less_le_trans";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   509
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   510
goal Nat.thy "!!i. [| i <= j; j <= k |] ==> i <= (k::nat)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   511
by (EVERY1[dtac le_imp_less_or_eq, dtac le_imp_less_or_eq,
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   512
          rtac less_or_eq_imp_le, fast_tac (!claset addIs [less_trans])]);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   513
qed "le_trans";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   514
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   515
val prems = goal Nat.thy "!!m. [| m <= n; n <= m |] ==> m = (n::nat)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   516
by (EVERY1[dtac le_imp_less_or_eq, dtac le_imp_less_or_eq,
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   517
          fast_tac (!claset addEs [less_irrefl,less_asym])]);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   518
qed "le_anti_sym";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   519
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   520
goal Nat.thy "(Suc(n) <= Suc(m)) = (n <= m)";
1264
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 1024
diff changeset
   521
by (simp_tac (!simpset addsimps [le_eq_less_or_eq]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   522
qed "Suc_le_mono";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   523
1264
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 1024
diff changeset
   524
Addsimps [le_refl,Suc_le_mono];
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   525
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   526
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   527
(** LEAST -- the least number operator **)
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   528
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   529
val [prem1,prem2] = goalw Nat.thy [Least_def]
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   530
    "[| P(k);  !!x. x<k ==> ~P(x) |] ==> (LEAST x.P(x)) = k";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   531
by (rtac select_equality 1);
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   532
by (fast_tac (!claset addSIs [prem1,prem2]) 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   533
by (cut_facts_tac [less_linear] 1);
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   534
by (fast_tac (!claset addSIs [prem1] addSDs [prem2]) 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   535
qed "Least_equality";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   536
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   537
val [prem] = goal Nat.thy "P(k) ==> P(LEAST x.P(x))";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   538
by (rtac (prem RS rev_mp) 1);
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   539
by (res_inst_tac [("n","k")] less_induct 1);
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   540
by (rtac impI 1);
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   541
by (rtac classical 1);
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   542
by (res_inst_tac [("s","n")] (Least_equality RS ssubst) 1);
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   543
by (assume_tac 1);
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   544
by (assume_tac 2);
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   545
by (Fast_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   546
qed "LeastI";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   547
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   548
(*Proof is almost identical to the one above!*)
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   549
val [prem] = goal Nat.thy "P(k) ==> (LEAST x.P(x)) <= k";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   550
by (rtac (prem RS rev_mp) 1);
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   551
by (res_inst_tac [("n","k")] less_induct 1);
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   552
by (rtac impI 1);
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   553
by (rtac classical 1);
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   554
by (res_inst_tac [("s","n")] (Least_equality RS ssubst) 1);
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   555
by (assume_tac 1);
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   556
by (rtac le_refl 2);
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   557
by (fast_tac (!claset addIs [less_imp_le,le_trans]) 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   558
qed "Least_le";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   559
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   560
val [prem] = goal Nat.thy "k < (LEAST x.P(x)) ==> ~P(k)";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   561
by (rtac notI 1);
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   562
by (etac (rewrite_rule [le_def] Least_le RS notE) 1);
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   563
by (rtac prem 1);
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1485
diff changeset
   564
qed "not_less_Least";
1660
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   565
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   566
qed_goalw "Least_Suc" Nat.thy [Least_def]
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   567
 "[| ? n. P (Suc n); ~ P 0 |] ==> (LEAST n. P n) = Suc (LEAST m. P (Suc m))"
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   568
 (fn prems => [
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   569
	cut_facts_tac prems 1,
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   570
	rtac select_equality 1,
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   571
	fold_goals_tac [Least_def],
1786
8a31d85d27b8 best_tac, deepen_tac and safe_tac now also use default claset.
berghofe
parents: 1777
diff changeset
   572
	safe_tac (!claset addSEs [LeastI]),
1660
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   573
	res_inst_tac [("n","j")] natE 1,
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   574
	Fast_tac 1,
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   575
	fast_tac (!claset addDs [Suc_less_SucD] addDs [not_less_Least]) 1,	
1660
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   576
	res_inst_tac [("n","k")] natE 1,
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   577
	Fast_tac 1,
1660
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   578
	hyp_subst_tac 1,
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   579
	rewtac Least_def,
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   580
	rtac (select_equality RS arg_cong RS sym) 1,
1786
8a31d85d27b8 best_tac, deepen_tac and safe_tac now also use default claset.
berghofe
parents: 1777
diff changeset
   581
	safe_tac (!claset),
1660
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   582
	dtac Suc_mono 1,
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   583
	Fast_tac 1,
1660
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   584
	cut_facts_tac [less_linear] 1,
1786
8a31d85d27b8 best_tac, deepen_tac and safe_tac now also use default claset.
berghofe
parents: 1777
diff changeset
   585
	safe_tac (!claset),
1660
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   586
	atac 2,
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   587
	Fast_tac 2,
1660
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   588
	dtac Suc_mono 1,
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1679
diff changeset
   589
	Fast_tac 1]);