| 5124 |      1 | (*  Title:      HOL/ex/Recdefs.ML
 | 
|  |      2 |     ID:         $Id$
 | 
| 8683 |      3 |     Author:     Konrad Slind and Lawrence C Paulson
 | 
| 5124 |      4 |     Copyright   1997  University of Cambridge
 | 
|  |      5 | 
 | 
|  |      6 | A few proofs to demonstrate the functions defined in Recdefs.thy
 | 
|  |      7 | Lemma statements from Konrad Slind's Web site
 | 
|  |      8 | *)
 | 
|  |      9 | 
 | 
|  |     10 | (** The silly g function: example of nested recursion **)
 | 
|  |     11 | 
 | 
| 8624 |     12 | Addsimps g.simps;
 | 
| 5124 |     13 | 
 | 
|  |     14 | Goal "g x < Suc x";
 | 
|  |     15 | by (res_inst_tac [("u","x")] g.induct 1);
 | 
|  |     16 | by Auto_tac;
 | 
|  |     17 | qed "g_terminates";
 | 
|  |     18 | 
 | 
|  |     19 | Goal "g x = 0";
 | 
|  |     20 | by (res_inst_tac [("u","x")] g.induct 1);
 | 
|  |     21 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [g_terminates])));
 | 
|  |     22 | qed "g_zero";
 | 
|  |     23 | 
 | 
|  |     24 | (*** the contrived `mapf' ***)
 | 
|  |     25 | 
 | 
|  |     26 | (* proving the termination condition: *)
 | 
|  |     27 | val [tc] = mapf.tcs;
 | 
|  |     28 | goalw_cterm [] (cterm_of (sign_of thy) (HOLogic.mk_Trueprop tc));
 | 
|  |     29 | by (rtac allI 1);
 | 
|  |     30 | by (case_tac "n=0" 1);
 | 
|  |     31 | by (ALLGOALS Asm_simp_tac);
 | 
|  |     32 | val lemma = result();
 | 
|  |     33 | 
 | 
|  |     34 | (* removing the termination condition from the generated thms: *)
 | 
| 8624 |     35 | val [mapf_0,mapf_Suc] = mapf.simps;
 | 
| 5124 |     36 | val mapf_Suc = lemma RS mapf_Suc;
 | 
|  |     37 | 
 | 
|  |     38 | val mapf_induct = lemma RS mapf.induct;
 |