Nominal/activities/tphols09/IDW/Inductive_Examples.thy
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Wed, 30 Mar 2016 17:27:34 +0100
changeset 415 f1be8028a4a9
permissions -rw-r--r--
updated

theory Inductive_Examples
imports Simple_Inductive_Package
begin

section {* Transitive closure *}

simple_inductive
  trcl for r :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
where
  base: "trcl r x x"
| step: "trcl r x y \<Longrightarrow> r y z \<Longrightarrow> trcl r x z"

thm trcl_def
thm trcl.induct
thm base
thm step
thm trcl.intros

lemma trcl_strong_induct:
  assumes trcl: "trcl r x y"
  and I1: "\<And>x. P x x"
  and I2: "\<And>x y z. P x y \<Longrightarrow> trcl r x y \<Longrightarrow> r y z \<Longrightarrow> P x z"
  shows "P x y" 
proof -
  from trcl
  have "P x y \<and> trcl r x y"
  proof induct
    case (base x)
    from I1 and trcl.base show ?case ..
  next
    case (step x y z)
    then have "P x y" and "trcl r x y" by simp_all
    from `P x y` `trcl r x y` `r y z` have "P x z"
      by (rule I2)
    moreover from `trcl r x y` `r y z` have "trcl r x z"
      by (rule trcl.step)
    ultimately show ?case ..
  qed
  then show ?thesis ..
qed


section {* Even and odd numbers *}

simple_inductive
  even and odd
where
  even0: "even 0"
| evenS: "odd n \<Longrightarrow> even (Suc n)"
| oddS: "even n \<Longrightarrow> odd (Suc n)"

thm even_def odd_def
thm even.induct odd.induct
thm even0
thm evenS
thm oddS
thm even_odd.intros


section {* Accessible part *}

simple_inductive
  accpart for r :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
where
  accpartI: "(\<And>y. r y x \<Longrightarrow> accpart r y) \<Longrightarrow> accpart r x"

thm accpart_def
thm accpart.induct
thm accpartI


section {* Accessible part in locale *}

locale rel =
  fixes r :: "'a \<Rightarrow> 'a \<Rightarrow> bool"

simple_inductive (in rel) accpart'
where
  accpartI': "\<And>x. (\<And>y. r y x \<Longrightarrow> accpart' y) \<Longrightarrow> accpart' x"

context rel
begin

thm accpartI'
thm accpart'.induct

end

thm rel.accpartI'
thm rel.accpart'.induct

end