src/LCF/ex/Ex4.thy
author wenzelm
Mon, 11 Feb 2013 14:39:04 +0100
changeset 51085 d90218288d51
parent 27208 5fe899199f85
child 58889 5b7a9633cfa8
permissions -rw-r--r--
make WWW_Find work again, now that its ML modules reside within a theory context (cf. bf5b45870110) -- patch by Rafal Kolanski;


header {* Prefixpoints *}

theory Ex4
imports LCF
begin

lemma example:
  assumes asms: "f(p) << p"  "!!q. f(q) << q ==> p << q"
  shows "FIX(f)=p"
  apply (unfold eq_def)
  apply (rule conjI)
  apply (tactic {* induct_tac @{context} "f" 1 *})
  apply (rule minimal)
  apply (intro strip)
  apply (rule less_trans)
  prefer 2
  apply (rule asms)
  apply (erule less_ap_term)
  apply (rule asms)
  apply (rule FIX_eq [THEN eq_imp_less1])
  done

end