Nominal/activities/cas09/Example.thy
author Christian Urban <urbanc@in.tum.de>
Fri, 19 Jul 2019 11:38:54 +0100
changeset 579 58c09367912e
parent 415 f1be8028a4a9
permissions -rw-r--r--
updated

theory Example
imports Main
begin

text {* 
  A small example theory proving a few facts
  about list append and list reverse.
*}

datatype 'a list2 =
    nil2                   ("[]")
  | cons2 "'a"  "'a list2" ("_ ## _")

abbreviation
  singleton :: "'a \<Rightarrow> 'a list2" ("[_]")
where
  "[x] \<equiv> x##[]"

fun
  append2 :: "'a list2 \<Rightarrow> 'a list2 \<Rightarrow> 'a list2" ("_ @@ _")
where
  append2_nil2:  "[] @@ ys = ys"
| append2_cons2: "(x##xs) @@ ys = x##(xs @@ ys)"

fun
  rev2 :: "'a list2 \<Rightarrow> 'a list2"
where
  "rev2 [] = []"
| "rev2 (x##xs) = (rev2 xs) @@ (x##[])"

lemma append2_nil2R[simp]:
  shows "xs @@ [] = xs"
by (induct xs) (auto)

lemma append2_assoc[simp]:
  shows "(xs @@ ys) @@ zs = xs @@ (ys @@ zs)"
by (induct xs) (auto)

lemma rev2_append2[simp]: 
  shows "rev2 (xs @@ ys) = (rev2 ys) @@ (rev2 xs)"
by (induct xs) (auto)

lemma rev2_rev2:
  shows "rev2 (rev2 xs) = xs"
by (induct xs) (auto)

end