| author | paulson | 
| Fri, 25 Sep 1998 14:05:13 +0200 | |
| changeset 5565 | 301a3a4d3dc7 | 
| parent 5316 | 7a8975451a89 | 
| child 5644 | 85fd64148873 | 
| permissions | -rw-r--r-- | 
| 2441 | 1 | (* Title: HOL/Nat.ML | 
| 923 | 2 | ID: $Id$ | 
| 2608 | 3 | Author: Tobias Nipkow | 
| 4 | Copyright 1997 TU Muenchen | |
| 923 | 5 | *) | 
| 6 | ||
| 5188 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 7 | (** conversion rules for nat_rec **) | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 8 | |
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 9 | val [nat_rec_0, nat_rec_Suc] = nat.recs; | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 10 | |
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 11 | (*These 2 rules ease the use of primitive recursion. NOTE USE OF == *) | 
| 5316 | 12 | val prems = Goal | 
| 5188 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 13 | "[| !!n. f(n) == nat_rec c h n |] ==> f(0) = c"; | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 14 | by (simp_tac (simpset() addsimps prems) 1); | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 15 | qed "def_nat_rec_0"; | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 16 | |
| 5316 | 17 | val prems = Goal | 
| 5188 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 18 | "[| !!n. f(n) == nat_rec c h n |] ==> f(Suc(n)) = h n (f n)"; | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 19 | by (simp_tac (simpset() addsimps prems) 1); | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 20 | qed "def_nat_rec_Suc"; | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 21 | |
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 22 | val [nat_case_0, nat_case_Suc] = nat.cases; | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 23 | |
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 24 | Goal "n ~= 0 ==> EX m. n = Suc m"; | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 25 | by (exhaust_tac "n" 1); | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 26 | by (REPEAT (Blast_tac 1)); | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 27 | qed "not0_implies_Suc"; | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 28 | |
| 5316 | 29 | Goal "m<n ==> n ~= 0"; | 
| 5188 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 30 | by (exhaust_tac "n" 1); | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 31 | by (ALLGOALS Asm_full_simp_tac); | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 32 | qed "gr_implies_not0"; | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 33 | |
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 34 | Goal "(n ~= 0) = (0 < n)"; | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 35 | by (exhaust_tac "n" 1); | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 36 | by (Blast_tac 1); | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 37 | by (Blast_tac 1); | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 38 | qed "neq0_conv"; | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 39 | AddIffs [neq0_conv]; | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 40 | |
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 41 | (*This theorem is useful with blast_tac: (n=0 ==> False) ==> 0<n *) | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 42 | bind_thm ("gr0I", [neq0_conv, notI] MRS iffD1);
 | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 43 | |
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 44 | Goal "(~(0 < n)) = (n=0)"; | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 45 | by (rtac iffI 1); | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 46 | by (etac swap 1); | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 47 | by (ALLGOALS Asm_full_simp_tac); | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 48 | qed "not_gr0"; | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 49 | Addsimps [not_gr0]; | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 50 | |
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 51 | Goal "m<n ==> 0 < n"; | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 52 | by (dtac gr_implies_not0 1); | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 53 | by (Asm_full_simp_tac 1); | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 54 | qed "gr_implies_gr0"; | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 55 | Addsimps [gr_implies_gr0]; | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 56 | |
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 57 | qed_goalw "Least_Suc" thy [Least_nat_def] | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 58 | "!!P. [| ? n. P(Suc n); ~ P 0 |] ==> (LEAST n. P n) = Suc (LEAST m. P(Suc m))" | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 59 | (fn _ => [ | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 60 | rtac select_equality 1, | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 61 | fold_goals_tac [Least_nat_def], | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 62 | safe_tac (claset() addSEs [LeastI]), | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 63 | rename_tac "j" 1, | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 64 | exhaust_tac "j" 1, | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 65 | Blast_tac 1, | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 66 | blast_tac (claset() addDs [Suc_less_SucD, not_less_Least]) 1, | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 67 | rename_tac "k n" 1, | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 68 | exhaust_tac "k" 1, | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 69 | Blast_tac 1, | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 70 | hyp_subst_tac 1, | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 71 | rewtac Least_nat_def, | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 72 | rtac (select_equality RS arg_cong RS sym) 1, | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 73 | Safe_tac, | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 74 | dtac Suc_mono 1, | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 75 | Blast_tac 1, | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 76 | cut_facts_tac [less_linear] 1, | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 77 | Safe_tac, | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 78 | atac 2, | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 79 | Blast_tac 2, | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 80 | dtac Suc_mono 1, | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 81 | Blast_tac 1]); | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 82 | |
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 83 | qed_goal "nat_induct2" thy | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 84 | "[| P 0; P 1; !!k. P k ==> P (Suc (Suc k)) |] ==> P n" (fn prems => [ | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 85 | cut_facts_tac prems 1, | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 86 | rtac less_induct 1, | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 87 | exhaust_tac "n" 1, | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 88 | hyp_subst_tac 1, | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 89 | atac 1, | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 90 | hyp_subst_tac 1, | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 91 | exhaust_tac "nat" 1, | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 92 | hyp_subst_tac 1, | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 93 | atac 1, | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 94 | hyp_subst_tac 1, | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 95 | resolve_tac prems 1, | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 96 | dtac spec 1, | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 97 | etac mp 1, | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 98 | rtac (lessI RS less_trans) 1, | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 99 | rtac (lessI RS Suc_mono) 1]); | 
| 
633ec5f6c155
Declaration of type 'nat' as a datatype (this allows usage of
 berghofe parents: 
5069diff
changeset | 100 | |
| 5069 | 101 | Goal "min 0 n = 0"; | 
| 3023 | 102 | by (rtac min_leastL 1); | 
| 103 | by (trans_tac 1); | |
| 2608 | 104 | qed "min_0L"; | 
| 1301 | 105 | |
| 5069 | 106 | Goal "min n 0 = 0"; | 
| 3023 | 107 | by (rtac min_leastR 1); | 
| 108 | by (trans_tac 1); | |
| 2608 | 109 | qed "min_0R"; | 
| 923 | 110 | |
| 5069 | 111 | Goalw [min_def] "min (Suc m) (Suc n) = Suc(min m n)"; | 
| 3023 | 112 | by (Simp_tac 1); | 
| 2608 | 113 | qed "min_Suc_Suc"; | 
| 1660 | 114 | |
| 2608 | 115 | Addsimps [min_0L,min_0R,min_Suc_Suc]; |