theory List imports Spartan begin (*TODO: Inductive type and recursive function definitions. The ad-hoc axiomatization below should be subsumed once general inductive types are properly implemented.*) axiomatization List :: \o \ o\ and nil :: \o \ o\ and cons :: \o \ o \ o \ o\ and ListInd :: \o \ (o \ o) \ o \ (o \ o \ o \ o) \ o \ o\ where ListF: "A: U i \ List A: U i" and List_nil: "A: U i \ nil A: List A" and List_cons: "\x: A; l: List A\ \ cons A x l: List A" and ListE: "\ xs: List A; c\<^sub>0: C (nil A); \x xs c. \x: A; xs: List A; c: C xs\ \ f x xs c: C (cons A x xs); \xs. xs: List A \ C xs: U i \ \ ListInd A (\xs. C xs) c\<^sub>0 (\x xs c. f x xs c) xs: C xs" and List_comp_nil: "\ c\<^sub>0: C (nil A); \x xs c. \x: A; xs: List A; c: C xs\ \ f x xs c: C (cons A x xs); \xs. xs: List A \ C xs: U i \ \ ListInd A (\xs. C xs) c\<^sub>0 (\x xs c. f x xs c) (nil A) \ c\<^sub>0" and List_comp_cons: "\ xs: List A; c\<^sub>0: C (nil A); \x xs c. \x: A; xs: List A; c: C xs\ \ f x xs c: C (cons A x xs); \xs. xs: List A \ C xs: U i \ \ ListInd A (\xs. C xs) c\<^sub>0 (\x xs c. f x xs c) (cons A x xs) \ f x xs (ListInd A (\xs. C xs) c\<^sub>0 (\x xs c. f x xs c) xs)" lemmas [intros] = ListF List_nil List_cons and [elims "?xs"] = ListE and [comps] = List_comp_nil List_comp_cons abbreviation "ListRec A C \ ListInd A (K C)" section \Implicit notation\ definition nil_i ("[]") where [implicit]: "[] \ nil ?" definition cons_i (infixr "#" 50) where [implicit]: "x # l \ cons ? x l" section \Standard functions\ definition "tail A \ \xs: List A. ListRec A (List A) (nil A) (\x xs _. xs) xs" definition "map A B \ \f: A \ B. \xs: List A. ListRec A (List B) (nil B) (\x _ c. cons B (f `x) c) xs" definition tail_i ("tail") where [implicit]: "tail xs \ List.tail ? xs" definition map_i ("map") where [implicit]: "map \ List.map ? ?" Lemma tail_type [typechk]: assumes "A: U i" "xs: List A" shows "tail xs: List A" unfolding tail_def by typechk Lemma map_type [typechk]: assumes "A: U i" "B: U i" "f: A \ B" "xs: List A" shows "map f xs: List B" unfolding map_def by typechk end