| author | blanchet |
| Mon, 21 Feb 2011 11:50:31 +0100 | |
| changeset 41793 | c7a2669ae75d |
| parent 35762 | af3ff2ba4c54 |
| child 47025 | b2b8ae61d6ad |
| permissions | -rw-r--r-- |
header {* Addition with fixpoint of successor *} theory Ex3 imports LCF begin consts s :: "'a => 'a" p :: "'a => 'a => 'a" axioms p_strict: "p(UU) = UU" p_s: "p(s(x),y) = s(p(x,y))" declare p_strict [simp] p_s [simp] lemma example: "p(FIX(s),y) = FIX(s)" apply (tactic {* induct_tac @{context} "s" 1 *}) apply (simp (no_asm)) apply (simp (no_asm)) done end