| author | wenzelm |
| Tue, 17 Mar 2015 15:21:41 +0100 | |
| changeset 59735 | 24bee1b11fce |
| parent 58977 | 9576b510f6a2 |
| child 60770 | 240563fbf41d |
| permissions | -rw-r--r-- |
section {* Addition with fixpoint of successor *} theory Ex3 imports "../LCF" begin axiomatization s :: "'a \<Rightarrow> 'a" and p :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" where p_strict: "p(UU) = UU" and 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 (induct s) apply simp apply simp done end