Nominal/activities/cas09/Example.thy
changeset 415 f1be8028a4a9
equal deleted inserted replaced
414:05e5d68c9627 415:f1be8028a4a9
       
     1 theory Example
       
     2 imports Main
       
     3 begin
       
     4 
       
     5 text {* 
       
     6   A small example theory proving a few facts
       
     7   about list append and list reverse.
       
     8 *}
       
     9 
       
    10 datatype 'a list2 =
       
    11     nil2                   ("[]")
       
    12   | cons2 "'a"  "'a list2" ("_ ## _")
       
    13 
       
    14 abbreviation
       
    15   singleton :: "'a \<Rightarrow> 'a list2" ("[_]")
       
    16 where
       
    17   "[x] \<equiv> x##[]"
       
    18 
       
    19 fun
       
    20   append2 :: "'a list2 \<Rightarrow> 'a list2 \<Rightarrow> 'a list2" ("_ @@ _")
       
    21 where
       
    22   append2_nil2:  "[] @@ ys = ys"
       
    23 | append2_cons2: "(x##xs) @@ ys = x##(xs @@ ys)"
       
    24 
       
    25 fun
       
    26   rev2 :: "'a list2 \<Rightarrow> 'a list2"
       
    27 where
       
    28   "rev2 [] = []"
       
    29 | "rev2 (x##xs) = (rev2 xs) @@ (x##[])"
       
    30 
       
    31 lemma append2_nil2R[simp]:
       
    32   shows "xs @@ [] = xs"
       
    33 by (induct xs) (auto)
       
    34 
       
    35 lemma append2_assoc[simp]:
       
    36   shows "(xs @@ ys) @@ zs = xs @@ (ys @@ zs)"
       
    37 by (induct xs) (auto)
       
    38 
       
    39 lemma rev2_append2[simp]: 
       
    40   shows "rev2 (xs @@ ys) = (rev2 ys) @@ (rev2 xs)"
       
    41 by (induct xs) (auto)
       
    42 
       
    43 lemma rev2_rev2:
       
    44   shows "rev2 (rev2 xs) = xs"
       
    45 by (induct xs) (auto)
       
    46 
       
    47 end