| 5124 |      1 | (*  Title:      HOL/ex/Recdefs.ML
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Konrad Lawrence C Paulson
 | 
|  |      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 | Addsimps qsort.rules;
 | 
|  |     11 | 
 | 
| 5518 |     12 | Goal "(x: set (qsort (ord,l))) = (x: set l)";
 | 
| 5124 |     13 | by (res_inst_tac [("u","ord"),("v","l")] qsort.induct 1);
 | 
| 6451 |     14 | by Auto_tac;
 | 
| 5124 |     15 | qed "qsort_mem_stable";
 | 
|  |     16 | 
 | 
|  |     17 | 
 | 
|  |     18 | (** The silly g function: example of nested recursion **)
 | 
|  |     19 | 
 | 
|  |     20 | Addsimps g.rules;
 | 
|  |     21 | 
 | 
|  |     22 | Goal "g x < Suc x";
 | 
|  |     23 | by (res_inst_tac [("u","x")] g.induct 1);
 | 
|  |     24 | by Auto_tac;
 | 
|  |     25 | qed "g_terminates";
 | 
|  |     26 | 
 | 
|  |     27 | Goal "g x = 0";
 | 
|  |     28 | by (res_inst_tac [("u","x")] g.induct 1);
 | 
|  |     29 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [g_terminates])));
 | 
|  |     30 | qed "g_zero";
 | 
|  |     31 | 
 | 
|  |     32 | (*** the contrived `mapf' ***)
 | 
|  |     33 | 
 | 
|  |     34 | (* proving the termination condition: *)
 | 
|  |     35 | val [tc] = mapf.tcs;
 | 
|  |     36 | goalw_cterm [] (cterm_of (sign_of thy) (HOLogic.mk_Trueprop tc));
 | 
|  |     37 | by (rtac allI 1);
 | 
|  |     38 | by (case_tac "n=0" 1);
 | 
|  |     39 | by (ALLGOALS Asm_simp_tac);
 | 
|  |     40 | val lemma = result();
 | 
|  |     41 | 
 | 
|  |     42 | (* removing the termination condition from the generated thms: *)
 | 
|  |     43 | val [mapf_0,mapf_Suc] = mapf.rules;
 | 
|  |     44 | val mapf_Suc = lemma RS mapf_Suc;
 | 
|  |     45 | 
 | 
|  |     46 | val mapf_induct = lemma RS mapf.induct;
 |