src/LCF/ex/Ex4.thy
author wenzelm
Sat, 23 May 2015 17:19:37 +0200
changeset 60299 5ae2a2e74c93
parent 58977 9576b510f6a2
child 60770 240563fbf41d
permissions -rw-r--r--
clarified NEWS: document_files are officially required since Isabelle2014, but the absence was tolerated as legacy feature;


section {* Prefixpoints *}

theory Ex4
imports "../LCF"
begin

lemma example:
  assumes asms: "f(p) << p"  "\<And>q. f(q) << q \<Longrightarrow> p << q"
  shows "FIX(f)=p"
  apply (unfold eq_def)
  apply (rule conjI)
  apply (induct f)
  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