src/HOL/ex/Seq.thy
author paulson <lp15@cam.ac.uk>
Fri, 23 Feb 2018 09:28:14 +0000
changeset 67691 db202a00a29c
parent 58889 5b7a9633cfa8
permissions -rw-r--r--
fixing ennreal using add_mono1; shifting results from linordered_semidom to linordered_nonzero_semiring

(*  Title:      HOL/ex/Seq.thy
    Author:     Makarius
*)

section \<open>Finite sequences\<close>

theory Seq
imports Main
begin

datatype 'a seq = Empty | Seq 'a "'a seq"

fun conc :: "'a seq \<Rightarrow> 'a seq \<Rightarrow> 'a seq"
where
  "conc Empty ys = ys"
| "conc (Seq x xs) ys = Seq x (conc xs ys)"

fun reverse :: "'a seq \<Rightarrow> 'a seq"
where
  "reverse Empty = Empty"
| "reverse (Seq x xs) = conc (reverse xs) (Seq x Empty)"

lemma conc_empty: "conc xs Empty = xs"
  by (induct xs) simp_all

lemma conc_assoc: "conc (conc xs ys) zs = conc xs (conc ys zs)"
  by (induct xs) simp_all

lemma reverse_conc: "reverse (conc xs ys) = conc (reverse ys) (reverse xs)"
  by (induct xs) (simp_all add: conc_empty conc_assoc)

lemma reverse_reverse: "reverse (reverse xs) = xs"
  by (induct xs) (simp_all add: reverse_conc)

end