Theory Tree_Forest

Up to index of Isabelle/ZF/Induct

theory Tree_Forest
imports Main
begin

(*  Title:      ZF/Induct/Tree_Forest.thy
    ID:         $Id: Tree_Forest.thy,v 1.11 2008/02/11 14:40:23 krauss Exp $
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
    Copyright   1994  University of Cambridge
*)

header {* Trees and forests, a mutually recursive type definition *}

theory Tree_Forest imports Main begin

subsection {* Datatype definition *}

consts
  tree :: "i => i"
  forest :: "i => i"
  tree_forest :: "i => i"

datatype "tree(A)" = Tcons ("a ∈ A", "f ∈ forest(A)")
  and "forest(A)" = Fnil | Fcons ("t ∈ tree(A)", "f ∈ forest(A)")

(* FIXME *)
lemmas tree'induct =
    tree_forest.mutual_induct [THEN conjunct1, THEN spec, THEN [2] rev_mp, of concl: _ t, standard, consumes 1]
  and forest'induct =
    tree_forest.mutual_induct [THEN conjunct2, THEN spec, THEN [2] rev_mp, of concl: _ f, standard, consumes 1]

declare tree_forest.intros [simp, TC]

lemma tree_def: "tree(A) == Part(tree_forest(A), Inl)"
  by (simp only: tree_forest.defs)

lemma forest_def: "forest(A) == Part(tree_forest(A), Inr)"
  by (simp only: tree_forest.defs)


text {*
  \medskip @{term "tree_forest(A)"} as the union of @{term "tree(A)"}
  and @{term "forest(A)"}.
*}

lemma tree_subset_TF: "tree(A) ⊆ tree_forest(A)"
  apply (unfold tree_forest.defs)
  apply (rule Part_subset)
  done

lemma treeI [TC]: "x ∈ tree(A) ==> x ∈ tree_forest(A)"
  by (rule tree_subset_TF [THEN subsetD])

lemma forest_subset_TF: "forest(A) ⊆ tree_forest(A)"
  apply (unfold tree_forest.defs)
  apply (rule Part_subset)
  done

lemma treeI' [TC]: "x ∈ forest(A) ==> x ∈ tree_forest(A)"
  by (rule forest_subset_TF [THEN subsetD])

lemma TF_equals_Un: "tree(A) ∪ forest(A) = tree_forest(A)"
  apply (insert tree_subset_TF forest_subset_TF)
  apply (auto intro!: equalityI tree_forest.intros elim: tree_forest.cases)
  done

lemma
  notes rews = tree_forest.con_defs tree_def forest_def
  shows
    tree_forest_unfold: "tree_forest(A) =
      (A × forest(A)) + ({0} + tree(A) × forest(A))"
    -- {* NOT useful, but interesting \dots *}
  apply (unfold tree_def forest_def)
  apply (fast intro!: tree_forest.intros [unfolded rews, THEN PartD1]
    elim: tree_forest.cases [unfolded rews])
  done

lemma tree_forest_unfold':
  "tree_forest(A) =
    A × Part(tree_forest(A), λw. Inr(w)) +
    {0} + Part(tree_forest(A), λw. Inl(w)) * Part(tree_forest(A), λw. Inr(w))"
  by (rule tree_forest_unfold [unfolded tree_def forest_def])

lemma tree_unfold: "tree(A) = {Inl(x). x ∈ A × forest(A)}"
  apply (unfold tree_def forest_def)
  apply (rule Part_Inl [THEN subst])
  apply (rule tree_forest_unfold' [THEN subst_context])
  done

lemma forest_unfold: "forest(A) = {Inr(x). x ∈ {0} + tree(A)*forest(A)}"
  apply (unfold tree_def forest_def)
  apply (rule Part_Inr [THEN subst])
  apply (rule tree_forest_unfold' [THEN subst_context])
  done

text {*
  \medskip Type checking for recursor: Not needed; possibly interesting?
*}

lemma TF_rec_type:
  "[| z ∈ tree_forest(A);
      !!x f r. [| x ∈ A;  f ∈ forest(A);  r ∈ C(f)
                |] ==> b(x,f,r) ∈ C(Tcons(x,f));
      c ∈ C(Fnil);
      !!t f r1 r2. [| t ∈ tree(A);  f ∈ forest(A);  r1 ∈ C(t); r2 ∈ C(f)
                    |] ==> d(t,f,r1,r2) ∈ C(Fcons(t,f))
   |] ==> tree_forest_rec(b,c,d,z) ∈ C(z)"
  by (induct_tac z) simp_all

lemma tree_forest_rec_type:
  "[| !!x f r. [| x ∈ A;  f ∈ forest(A);  r ∈ D(f)
                |] ==> b(x,f,r) ∈ C(Tcons(x,f));
      c ∈ D(Fnil);
      !!t f r1 r2. [| t ∈ tree(A);  f ∈ forest(A);  r1 ∈ C(t); r2 ∈ D(f)
                    |] ==> d(t,f,r1,r2) ∈ D(Fcons(t,f))
   |] ==> (∀t ∈ tree(A).    tree_forest_rec(b,c,d,t) ∈ C(t)) ∧
          (∀f ∈ forest(A). tree_forest_rec(b,c,d,f) ∈ D(f))"
    -- {* Mutually recursive version. *}
  apply (unfold Ball_def)
  apply (rule tree_forest.mutual_induct)
  apply simp_all
  done


subsection {* Operations *}

consts
  map :: "[i => i, i] => i"
  size :: "i => i"
  preorder :: "i => i"
  list_of_TF :: "i => i"
  of_list :: "i => i"
  reflect :: "i => i"

primrec
  "list_of_TF (Tcons(x,f)) = [Tcons(x,f)]"
  "list_of_TF (Fnil) = []"
  "list_of_TF (Fcons(t,tf)) = Cons (t, list_of_TF(tf))"

primrec
  "of_list([]) = Fnil"
  "of_list(Cons(t,l)) = Fcons(t, of_list(l))"

primrec
  "map (h, Tcons(x,f)) = Tcons(h(x), map(h,f))"
  "map (h, Fnil) = Fnil"
  "map (h, Fcons(t,tf)) = Fcons (map(h, t), map(h, tf))"

primrec
  "size (Tcons(x,f)) = succ(size(f))"
  "size (Fnil) = 0"
  "size (Fcons(t,tf)) = size(t) #+ size(tf)"

primrec
  "preorder (Tcons(x,f)) = Cons(x, preorder(f))"
  "preorder (Fnil) = Nil"
  "preorder (Fcons(t,tf)) = preorder(t) @ preorder(tf)"

primrec
  "reflect (Tcons(x,f)) = Tcons(x, reflect(f))"
  "reflect (Fnil) = Fnil"
  "reflect (Fcons(t,tf)) =
    of_list (list_of_TF (reflect(tf)) @ Cons(reflect(t), Nil))"


text {*
  \medskip @{text list_of_TF} and @{text of_list}.
*}

lemma list_of_TF_type [TC]:
    "z ∈ tree_forest(A) ==> list_of_TF(z) ∈ list(tree(A))"
  by (induct set: tree_forest) simp_all

lemma of_list_type [TC]: "l ∈ list(tree(A)) ==> of_list(l) ∈ forest(A)"
  by (induct set: list) simp_all

text {*
  \medskip @{text map}.
*}

lemma
  assumes "!!x. x ∈ A ==> h(x): B"
  shows map_tree_type: "t ∈ tree(A) ==> map(h,t) ∈ tree(B)"
    and map_forest_type: "f ∈ forest(A) ==> map(h,f) ∈ forest(B)"
  using prems
  by (induct rule: tree'induct forest'induct) simp_all

text {*
  \medskip @{text size}.
*}

lemma size_type [TC]: "z ∈ tree_forest(A) ==> size(z) ∈ nat"
  by (induct set: tree_forest) simp_all


text {*
  \medskip @{text preorder}.
*}

lemma preorder_type [TC]: "z ∈ tree_forest(A) ==> preorder(z) ∈ list(A)"
  by (induct set: tree_forest) simp_all


text {*
  \medskip Theorems about @{text list_of_TF} and @{text of_list}.
*}

lemma forest_induct [consumes 1, case_names Fnil Fcons]:
  "[| f ∈ forest(A);
      R(Fnil);
      !!t f. [| t ∈ tree(A);  f ∈ forest(A);  R(f) |] ==> R(Fcons(t,f))
   |] ==> R(f)"
  -- {* Essentially the same as list induction. *}
  apply (erule tree_forest.mutual_induct
      [THEN conjunct2, THEN spec, THEN [2] rev_mp])
    apply (rule TrueI)
   apply simp
  apply simp
  done

lemma forest_iso: "f ∈ forest(A) ==> of_list(list_of_TF(f)) = f"
  by (induct rule: forest_induct) simp_all

lemma tree_list_iso: "ts: list(tree(A)) ==> list_of_TF(of_list(ts)) = ts"
  by (induct set: list) simp_all


text {*
  \medskip Theorems about @{text map}.
*}

lemma map_ident: "z ∈ tree_forest(A) ==> map(λu. u, z) = z"
  by (induct set: tree_forest) simp_all

lemma map_compose:
    "z ∈ tree_forest(A) ==> map(h, map(j,z)) = map(λu. h(j(u)), z)"
  by (induct set: tree_forest) simp_all


text {*
  \medskip Theorems about @{text size}.
*}

lemma size_map: "z ∈ tree_forest(A) ==> size(map(h,z)) = size(z)"
  by (induct set: tree_forest) simp_all

lemma size_length: "z ∈ tree_forest(A) ==> size(z) = length(preorder(z))"
  by (induct set: tree_forest) (simp_all add: length_app)

text {*
  \medskip Theorems about @{text preorder}.
*}

lemma preorder_map:
    "z ∈ tree_forest(A) ==> preorder(map(h,z)) = List_ZF.map(h, preorder(z))"
  by (induct set: tree_forest) (simp_all add: map_app_distrib)

end

Datatype definition

lemma tree'induct:

  [| t ∈ tree(A);
     !!a f. [| aA; f ∈ forest(A); P_forest(f) |] ==> P_tree(Tcons(a, f));
     P_forest(Fnil);
     !!f t. [| t ∈ tree(A); P_tree(t); f ∈ forest(A); P_forest(f) |]
            ==> P_forest(Fcons(t, f)) |]
  ==> P_tree(t)

and forest'induct:

  [| f ∈ forest(A);
     !!a f. [| aA; f ∈ forest(A); P_forest(f) |] ==> P_tree(Tcons(a, f));
     P_forest(Fnil);
     !!f t. [| t ∈ tree(A); P_tree(t); f ∈ forest(A); P_forest(f) |]
            ==> P_forest(Fcons(t, f)) |]
  ==> P_forest(f)

lemma tree_def:

  tree(A) == Part(tree_forest(A), Inl)

lemma forest_def:

  forest(A) == Part(tree_forest(A), Inr)

lemma tree_subset_TF:

  tree(A) ⊆ tree_forest(A)

lemma treeI:

  x ∈ tree(A) ==> x ∈ tree_forest(A)

lemma forest_subset_TF:

  forest(A) ⊆ tree_forest(A)

lemma treeI':

  x ∈ forest(A) ==> x ∈ tree_forest(A)

lemma TF_equals_Un:

  tree(A) ∪ forest(A) = tree_forest(A)

lemma tree_forest_unfold:

  tree_forest(A) = A × forest(A) + {0} + tree(A) × forest(A)

lemma tree_forest_unfold':

  tree_forest(A) =
  A × Part(tree_forest(A), λw. Inr(w)) +
  {0} + Part(tree_forest(A), λw. Inl(w)) × Part(tree_forest(A), λw. Inr(w))

lemma tree_unfold:

  tree(A) = {Inl(x) . xA × forest(A)}

lemma forest_unfold:

  forest(A) = {Inr(x) . x ∈ {0} + tree(A) × forest(A)}

lemma TF_rec_type:

  [| z ∈ tree_forest(A);
     !!x f r.
        [| xA; f ∈ forest(A); rC(f) |] ==> b(x, f, r) ∈ C(Tcons(x, f));
     cC(Fnil);
     !!t f r1 r2.
        [| t ∈ tree(A); f ∈ forest(A); r1C(t); r2C(f) |]
        ==> d(t, f, r1, r2) ∈ C(Fcons(t, f)) |]
  ==> tree_forest_rec(b, c, d, z) ∈ C(z)

lemma tree_forest_rec_type:

  [| !!x f r.
        [| xA; f ∈ forest(A); rD(f) |] ==> b(x, f, r) ∈ C(Tcons(x, f));
     cD(Fnil);
     !!t f r1 r2.
        [| t ∈ tree(A); f ∈ forest(A); r1C(t); r2D(f) |]
        ==> d(t, f, r1, r2) ∈ D(Fcons(t, f)) |]
  ==> (∀t∈tree(A). tree_forest_rec(b, c, d, t) ∈ C(t)) ∧
      (∀f∈forest(A). tree_forest_rec(b, c, d, f) ∈ D(f))

Operations

lemma list_of_TF_type:

  z ∈ tree_forest(A) ==> list_of_TF(z) ∈ list(tree(A))

lemma of_list_type:

  l ∈ list(tree(A)) ==> of_list(l) ∈ forest(A)

lemma map_tree_type:

  [| !!x. xA ==> h(x) ∈ B; t ∈ tree(A) |] ==> Tree_Forest.map(h, t) ∈ tree(B)

and map_forest_type:

  [| !!x. xA ==> h(x) ∈ B; f ∈ forest(A) |]
  ==> Tree_Forest.map(h, f) ∈ forest(B)

lemma size_type:

  z ∈ tree_forest(A) ==> size(z) ∈ nat

lemma preorder_type:

  z ∈ tree_forest(A) ==> preorder(z) ∈ list(A)

lemma forest_induct:

  [| f ∈ forest(A); R(Fnil);
     !!t f. [| t ∈ tree(A); f ∈ forest(A); R(f) |] ==> R(Fcons(t, f)) |]
  ==> R(f)

lemma forest_iso:

  f ∈ forest(A) ==> of_list(list_of_TF(f)) = f

lemma tree_list_iso:

  ts ∈ list(tree(A)) ==> list_of_TF(of_list(ts)) = ts

lemma map_ident:

  z ∈ tree_forest(A) ==> Tree_Forest.map(λu. u, z) = z

lemma map_compose:

  z ∈ tree_forest(A)
  ==> Tree_Forest.map(h, Tree_Forest.map(j, z)) = Tree_Forest.map(λu. h(j(u)), z)

lemma size_map:

  z ∈ tree_forest(A) ==> size(Tree_Forest.map(h, z)) = size(z)

lemma size_length:

  z ∈ tree_forest(A) ==> size(z) = length(preorder(z))

lemma preorder_map:

  z ∈ tree_forest(A)
  ==> preorder(Tree_Forest.map(h, z)) = List_ZF.map(h, preorder(z))