| author | hoelzl | 
| Tue, 24 Aug 2010 14:41:37 +0200 | |
| changeset 38705 | aaee86c0e237 | 
| parent 35762 | af3ff2ba4c54 | 
| child 47025 | b2b8ae61d6ad | 
| permissions | -rw-r--r-- | 
| 17248 | 1 | header {* Addition with fixpoint of successor *}
 | 
| 4905 | 2 | |
| 17248 | 3 | theory Ex3 | 
| 4 | imports LCF | |
| 5 | begin | |
| 4905 | 6 | |
| 7 | consts | |
| 17248 | 8 | s :: "'a => 'a" | 
| 9 | p :: "'a => 'a => 'a" | |
| 4905 | 10 | |
| 17248 | 11 | axioms | 
| 12 | p_strict: "p(UU) = UU" | |
| 13 | p_s: "p(s(x),y) = s(p(x,y))" | |
| 14 | ||
| 19755 | 15 | declare p_strict [simp] p_s [simp] | 
| 16 | ||
| 17 | lemma example: "p(FIX(s),y) = FIX(s)" | |
| 27208 
5fe899199f85
proper context for tactics derived from res_inst_tac;
 wenzelm parents: 
19755diff
changeset | 18 |   apply (tactic {* induct_tac @{context} "s" 1 *})
 | 
| 19755 | 19 | apply (simp (no_asm)) | 
| 20 | apply (simp (no_asm)) | |
| 21 | done | |
| 4905 | 22 | |
| 23 | end |