| author | berghofe | 
| Mon, 07 Nov 2005 18:32:54 +0100 | |
| changeset 18107 | ee6b4d3af498 | 
| parent 17456 | bcf7544875b2 | 
| permissions | -rw-r--r-- | 
| 17456 | 1 | (* Title: CCL/ex/Nat.ML | 
| 0 | 2 | ID: $Id$ | 
| 1459 | 3 | Author: Martin Coen, Cambridge University Computer Laboratory | 
| 0 | 4 | Copyright 1993 University of Cambridge | 
| 5 | *) | |
| 6 | ||
| 7 | val nat_defs = [not_def,add_def,mult_def,sub_def,le_def,lt_def,ack_def,napply_def]; | |
| 8 | ||
| 17456 | 9 | val natBs = map (fn s=>prove_goalw (the_context ()) nat_defs s (fn _ => [simp_tac term_ss 1])) | 
| 0 | 10 | ["not(true) = false", | 
| 11 | "not(false) = true", | |
| 12 | "zero #+ n = n", | |
| 13 | "succ(n) #+ m = succ(n #+ m)", | |
| 14 | "zero #* n = zero", | |
| 15 | "succ(n) #* m = m #+ (n #* m)", | |
| 16 | "f^zero`a = a", | |
| 17 | "f^succ(n)`a = f(f^n`a)"]; | |
| 18 | ||
| 8 
c3d2c6dcf3f0
Installation of new simplfier.  Previously appeared to set up the old
 lcp parents: 
0diff
changeset | 19 | val nat_ss = term_ss addsimps natBs; | 
| 0 | 20 | |
| 21 | (*** Lemma for napply ***) | |
| 22 | ||
| 17456 | 23 | val [prem] = goal (the_context ()) "n:Nat ==> f^n`f(a) = f^succ(n)`a"; | 
| 1459 | 24 | by (rtac (prem RS Nat_ind) 1); | 
| 8 
c3d2c6dcf3f0
Installation of new simplfier.  Previously appeared to set up the old
 lcp parents: 
0diff
changeset | 25 | by (ALLGOALS (asm_simp_tac nat_ss)); | 
| 757 | 26 | qed "napply_f"; | 
| 0 | 27 | |
| 28 | (****) | |
| 29 | ||
| 17456 | 30 | val prems = goalw (the_context ()) [add_def] "[| a:Nat; b:Nat |] ==> a #+ b : Nat"; | 
| 0 | 31 | by (typechk_tac prems 1); | 
| 757 | 32 | qed "addT"; | 
| 0 | 33 | |
| 17456 | 34 | val prems = goalw (the_context ()) [mult_def] "[| a:Nat; b:Nat |] ==> a #* b : Nat"; | 
| 0 | 35 | by (typechk_tac (addT::prems) 1); | 
| 757 | 36 | qed "multT"; | 
| 0 | 37 | |
| 38 | (* Defined to return zero if a<b *) | |
| 17456 | 39 | val prems = goalw (the_context ()) [sub_def] "[| a:Nat; b:Nat |] ==> a #- b : Nat"; | 
| 0 | 40 | by (typechk_tac (prems) 1); | 
| 41 | by clean_ccs_tac; | |
| 1459 | 42 | by (etac (NatPRI RS wfstI RS (NatPR_wf RS wmap_wf RS wfI)) 1); | 
| 757 | 43 | qed "subT"; | 
| 0 | 44 | |
| 17456 | 45 | val prems = goalw (the_context ()) [le_def] "[| a:Nat; b:Nat |] ==> a #<= b : Bool"; | 
| 0 | 46 | by (typechk_tac (prems) 1); | 
| 47 | by clean_ccs_tac; | |
| 1459 | 48 | by (etac (NatPRI RS wfstI RS (NatPR_wf RS wmap_wf RS wfI)) 1); | 
| 757 | 49 | qed "leT"; | 
| 0 | 50 | |
| 17456 | 51 | val prems = goalw (the_context ()) [not_def,lt_def] "[| a:Nat; b:Nat |] ==> a #< b : Bool"; | 
| 0 | 52 | by (typechk_tac (prems@[leT]) 1); | 
| 757 | 53 | qed "ltT"; | 
| 0 | 54 | |
| 55 | (* Correctness conditions for subtractive division **) | |
| 56 | ||
| 17456 | 57 | val prems = goalw (the_context ()) [div_def] | 
| 0 | 58 |     "[| a:Nat;  b:{x:Nat.~x=zero} |] ==> a ## b : {x:Nat. DIV(a,b,x)}";
 | 
| 59 | by (gen_ccs_tac (prems@[ltT,subT]) 1); | |
| 60 | ||
| 61 | (* Termination Conditions for Ackermann's Function *) | |
| 62 | ||
| 17456 | 63 | val prems = goalw (the_context ()) [ack_def] | 
| 0 | 64 | "[| a:Nat; b:Nat |] ==> ackermann(a,b) : Nat"; | 
| 65 | by (gen_ccs_tac prems 1); | |
| 66 | val relI = NatPR_wf RS (NatPR_wf RS lex_wf RS wfI); | |
| 67 | by (REPEAT (eresolve_tac [NatPRI RS (lexI1 RS relI),NatPRI RS (lexI2 RS relI)] 1)); | |
| 68 | result(); |