| 60770 |      1 | section \<open>Section 10.4\<close>
 | 
| 4905 |      2 | 
 | 
| 17248 |      3 | theory Ex1
 | 
| 58974 |      4 | imports "../LCF"
 | 
| 17248 |      5 | begin
 | 
| 4905 |      6 | 
 | 
| 47025 |      7 | axiomatization
 | 
| 58977 |      8 |   P     :: "'a \<Rightarrow> tr" and
 | 
|  |      9 |   G     :: "'a \<Rightarrow> 'a" and
 | 
|  |     10 |   H     :: "'a \<Rightarrow> 'a" and
 | 
|  |     11 |   K     :: "('a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a)"
 | 
| 47025 |     12 | where
 | 
|  |     13 |   P_strict:     "P(UU) = UU" and
 | 
| 58977 |     14 |   K:            "K = (\<lambda>h x. P(x) \<Rightarrow> x | h(h(G(x))))" and
 | 
| 17248 |     15 |   H:            "H = FIX(K)"
 | 
|  |     16 | 
 | 
| 19755 |     17 | 
 | 
|  |     18 | declare P_strict [simp] K [simp]
 | 
|  |     19 | 
 | 
|  |     20 | lemma H_unfold: "H = K(H)"
 | 
|  |     21 |   apply (simplesubst H)
 | 
|  |     22 |   apply (rule FIX_eq [symmetric])
 | 
|  |     23 |   done
 | 
|  |     24 | 
 | 
|  |     25 | lemma H_strict [simp]: "H(UU)=UU"
 | 
|  |     26 |   apply (simplesubst H_unfold)
 | 
|  |     27 |   apply simp
 | 
|  |     28 |   done
 | 
|  |     29 | 
 | 
| 58977 |     30 | lemma H_idemp_lemma: "\<forall>x. H(FIX(K,x)) = FIX(K,x)"
 | 
| 58973 |     31 |   apply (induct K)
 | 
| 47025 |     32 |   apply simp
 | 
|  |     33 |   apply (simp split: COND_cases_iff)
 | 
| 19755 |     34 |   apply (intro strip)
 | 
|  |     35 |   apply (subst H_unfold)
 | 
| 47025 |     36 |   apply simp
 | 
| 19755 |     37 |   done
 | 
|  |     38 | 
 | 
| 58977 |     39 | lemma H_idemp: "\<forall>x. H(H(x)) = H(x)"
 | 
| 19755 |     40 |   apply (rule H_idemp_lemma [folded H])
 | 
|  |     41 |   done
 | 
| 4905 |     42 | 
 | 
|  |     43 | end
 |