adamc@193: (* Copyright (c) 2008, Adam Chlipala adamc@193: * adamc@193: * This work is licensed under a adamc@193: * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0 adamc@193: * Unported License. adamc@193: * The license text is available at: adamc@193: * http://creativecommons.org/licenses/by-nc-nd/3.0/ adamc@193: *) adamc@193: adamc@193: (* begin hide *) adamc@195: Require Import String List. adamc@193: adamc@193: Require Import Tactics DepList. adamc@193: adamc@193: Set Implicit Arguments. adamc@193: (* end hide *) adamc@193: adamc@193: adamc@193: (** %\part{Chapters to be Moved Earlier} adamc@193: adamc@193: \chapter{Generic Programming}% *) adamc@193: adamc@193: (** TODO: Prose for this chapter *) adamc@193: adamc@193: adamc@195: (** * Reflecting Datatype Definitions *) adamc@193: adamc@193: Record constructor : Type := Con { adamc@193: nonrecursive : Type; adamc@193: recursive : nat adamc@193: }. adamc@193: adamc@193: Definition datatype := list constructor. adamc@193: adamc@193: Definition Empty_set_dt : datatype := nil. adamc@193: Definition unit_dt : datatype := Con unit 0 :: nil. adamc@193: Definition bool_dt : datatype := Con unit 0 :: Con unit 0 :: nil. adamc@193: Definition nat_dt : datatype := Con unit 0 :: Con unit 1 :: nil. adamc@193: Definition list_dt (A : Type) : datatype := Con unit 0 :: Con A 1 :: nil. adamc@193: adamc@193: Section tree. adamc@193: Variable A : Type. adamc@193: adamc@193: Inductive tree : Type := adamc@193: | Leaf : A -> tree adamc@193: | Node : tree -> tree -> tree. adamc@193: End tree. adamc@193: adamc@193: Definition tree_dt (A : Type) : datatype := Con A 0 :: Con unit 2 :: nil. adamc@193: adamc@193: Section denote. adamc@193: Variable T : Type. adamc@193: adamc@193: Definition constructorDenote (c : constructor) := adamc@193: nonrecursive c -> ilist T (recursive c) -> T. adamc@193: adamc@193: Definition datatypeDenote := hlist constructorDenote. adamc@193: End denote. adamc@193: adamc@193: Notation "[ ! , ! ~> x ]" := ((fun _ _ => x) : constructorDenote _ (Con _ _)). adamc@193: Notation "[ v , ! ~> x ]" := ((fun v _ => x) : constructorDenote _ (Con _ _)). adamc@193: Notation "[ ! , r # n ~> x ]" := ((fun _ r => x) : constructorDenote _ (Con _ n)). adamc@193: Notation "[ v , r # n ~> x ]" := ((fun v r => x) : constructorDenote _ (Con _ n)). adamc@193: adamc@193: Definition Empty_set_den : datatypeDenote Empty_set Empty_set_dt := adamc@193: hnil. adamc@193: Definition unit_den : datatypeDenote unit unit_dt := adamc@193: [!, ! ~> tt] ::: hnil. adamc@193: Definition bool_den : datatypeDenote bool bool_dt := adamc@193: [!, ! ~> true] ::: [!, ! ~> false] ::: hnil. adamc@193: Definition nat_den : datatypeDenote nat nat_dt := adamc@193: [!, ! ~> O] ::: [!, r # 1 ~> S (hd r)] ::: hnil. adamc@193: Definition list_den (A : Type) : datatypeDenote (list A) (list_dt A) := adamc@193: [!, ! ~> nil] ::: [x, r # 1 ~> x :: hd r] ::: hnil. adamc@193: Definition tree_den (A : Type) : datatypeDenote (tree A) (tree_dt A) := adamc@193: [v, ! ~> Leaf v] ::: [!, r # 2 ~> Node (hd r) (hd (tl r))] ::: hnil. adamc@194: adamc@195: adamc@195: (** * Recursive Definitions *) adamc@195: adamc@194: Definition fixDenote (T : Type) (dt : datatype) := adamc@194: forall (R : Type), datatypeDenote R dt -> (T -> R). adamc@194: adamc@194: Definition size T dt (fx : fixDenote T dt) : T -> nat := adamc@194: fx nat (hmake (B := constructorDenote nat) (fun _ _ r => foldr plus 1 r) dt). adamc@194: adamc@194: Definition Empty_set_fix : fixDenote Empty_set Empty_set_dt := adamc@194: fun R _ emp => match emp with end. adamc@194: Eval compute in size Empty_set_fix. adamc@194: adamc@194: Definition unit_fix : fixDenote unit unit_dt := adamc@194: fun R cases _ => (fst cases) tt inil. adamc@194: Eval compute in size unit_fix. adamc@194: adamc@194: Definition bool_fix : fixDenote bool bool_dt := adamc@194: fun R cases b => if b adamc@194: then (fst cases) tt inil adamc@194: else (fst (snd cases)) tt inil. adamc@194: Eval compute in size bool_fix. adamc@194: adamc@194: Definition nat_fix : fixDenote nat nat_dt := adamc@194: fun R cases => fix F (n : nat) : R := adamc@194: match n with adamc@194: | O => (fst cases) tt inil adamc@194: | S n' => (fst (snd cases)) tt (icons (F n') inil) adamc@194: end. adamc@194: Eval cbv beta iota delta -[plus] in size nat_fix. adamc@194: adamc@194: Definition list_fix (A : Type) : fixDenote (list A) (list_dt A) := adamc@194: fun R cases => fix F (ls : list A) : R := adamc@194: match ls with adamc@194: | nil => (fst cases) tt inil adamc@194: | x :: ls' => (fst (snd cases)) x (icons (F ls') inil) adamc@194: end. adamc@194: Eval cbv beta iota delta -[plus] in fun A => size (@list_fix A). adamc@194: adamc@194: Definition tree_fix (A : Type) : fixDenote (tree A) (tree_dt A) := adamc@194: fun R cases => fix F (t : tree A) : R := adamc@194: match t with adamc@194: | Leaf x => (fst cases) x inil adamc@194: | Node t1 t2 => (fst (snd cases)) tt (icons (F t1) (icons (F t2) inil)) adamc@194: end. adamc@194: Eval cbv beta iota delta -[plus] in fun A => size (@tree_fix A). adamc@195: adamc@195: adamc@195: (** ** Pretty-Printing *) adamc@195: adamc@195: Record print_constructor (c : constructor) : Type := PI { adamc@195: printName : string; adamc@195: printNonrec : nonrecursive c -> string adamc@195: }. adamc@195: adamc@195: Notation "^" := (PI (Con _ _)). adamc@195: adamc@195: Definition print_datatype := hlist print_constructor. adamc@195: adamc@195: Open Local Scope string_scope. adamc@195: adamc@195: Definition print T dt (pr : print_datatype dt) (fx : fixDenote T dt) : T -> string := adamc@195: fx string (hmap (B1 := print_constructor) (B2 := constructorDenote string) adamc@195: (fun _ pc x r => printName pc ++ "(" ++ printNonrec pc x adamc@195: ++ foldr (fun s acc => ", " ++ s ++ acc) ")" r) pr). adamc@195: adamc@195: Eval compute in print hnil Empty_set_fix. adamc@195: Eval compute in print (^ "tt" (fun _ => "") ::: hnil) unit_fix. adamc@195: Eval compute in print (^ "true" (fun _ => "") adamc@195: ::: ^ "false" (fun _ => "") adamc@195: ::: hnil) bool_fix. adamc@195: adamc@195: Definition print_nat := print (^ "O" (fun _ => "") adamc@195: ::: ^ "S" (fun _ => "") adamc@195: ::: hnil) nat_fix. adamc@195: Eval cbv beta iota delta -[append] in print_nat. adamc@195: Eval simpl in print_nat 0. adamc@195: Eval simpl in print_nat 1. adamc@195: Eval simpl in print_nat 2. adamc@195: adamc@195: Eval cbv beta iota delta -[append] in fun A (pr : A -> string) => adamc@195: print (^ "nil" (fun _ => "") adamc@195: ::: ^ "cons" pr adamc@195: ::: hnil) (@list_fix A). adamc@195: Eval cbv beta iota delta -[append] in fun A (pr : A -> string) => adamc@195: print (^ "Leaf" pr adamc@195: ::: ^ "Node" (fun _ => "") adamc@195: ::: hnil) (@tree_fix A).