(* week-01_polymorphism.v *) (* LPP 2024 - CS3234 2023-2024, Sem2 *) (* Olivier Danvy *) (* Version of 18 Jan 2024 *) (* ********** *) Require Import Arith Bool. (* ********** *) Definition make_pair_v0 : forall (A : Type) (B : Type), A -> B -> A * B := fun A B a b => (a, b). Definition make_pair_v1 (A : Type) : forall B : Type, A -> B -> A * B := fun B a b => (a, b). Definition make_pair_v2 (A : Type) (B : Type) : A -> B -> A * B := fun a b => (a, b). Definition make_pair_v3 (A : Type) (B : Type) (a : A) : B -> A * B := fun b => (a, b). Definition make_pair_v4 (A : Type) (B : Type) (a : A) (b : B) : A * B := (a, b). (* ***** *) Compute make_pair_v0 nat bool 1 true. Compute (make_pair_v4 nat bool 1 true). (* = (1, true) : nat * bool *) Definition dupl (V : Type) (v : V) : V * V := make_pair_v4 V V v v. (* ********** *) Inductive polymorphic_binary_tree (V : Type) : Type := | PLeaf : V -> polymorphic_binary_tree V | PNode : polymorphic_binary_tree V -> polymorphic_binary_tree V -> polymorphic_binary_tree V. Definition test_number_of_leaves (candidate : forall V : Type, polymorphic_binary_tree V -> nat) := (candidate bool (PLeaf bool true) =? 1) && (candidate nat (PNode nat (PLeaf nat 0) (PLeaf nat 1)) =? 2). Fixpoint number_of_leaves (V : Type) (t : polymorphic_binary_tree V) : nat := match t with | PLeaf _ v => 1 | PNode _ t1 t2 => number_of_leaves V t1 + number_of_leaves V t2 end. Compute (test_number_of_leaves number_of_leaves). (* ********** *) Fixpoint eqb_polymorphic_binary_tree (V : Type) (eqb_V : V -> V -> bool) (t1 t2 : polymorphic_binary_tree V) : bool := match t1 with | PLeaf _ v1 => match t2 with | PLeaf _ v2 => eqb_V v1 v2 | PNode _ t11 t12 => false end | PNode _ t11 t12 => match t2 with | PLeaf _ v2 => false | PNode _ t21 t22 => eqb_polymorphic_binary_tree V eqb_V t11 t21 && eqb_polymorphic_binary_tree V eqb_V t12 t22 end end. (* ********** *) Definition eqb_nat := Nat.eqb. Definition eqb_binary_tree_of_nats (t1 t2 : polymorphic_binary_tree nat) : bool := eqb_polymorphic_binary_tree nat eqb_nat t1 t2. (* ********** *) (* Exercise 05 *) Definition e5a : polymorphic_binary_tree (nat * bool) := PNode (nat * bool) (PLeaf (nat * bool) (0, true)) (PLeaf (nat * bool) (1, false)) . Check e5a. Compute e5a. Definition nTree := (polymorphic_binary_tree nat). Definition e5b : polymorphic_binary_tree (polymorphic_binary_tree nat) := PNode (nTree) (PLeaf (nTree) (PLeaf nat 1)) (PLeaf (nTree) (PLeaf nat 2)) . Check e5b. Compute e5b. (* Exercise 07 *) Definition eqb_option (V : Type) (eqb_V : V -> V -> bool) (ov1 ov2 : option V) : bool := match ov1 with | Some v1 => match ov2 with | Some v2 => eqb_V v1 v2 | None => false end | None => match ov2 with | Some v => false | None => true end end. Check (eqb_polymorphic_binary_tree nat). Check (eqb_polymorphic_binary_tree (option nat)). Check (eqb_option (polymorphic_binary_tree nat)). (* Definition eqb_binary_tree_of_optional_nats (t1 t2 : polymorphic_binary_tree (option nat)) : bool := *) (* Definition eqb_optional_binary_tree_of_nats (t1 t2 : option (polymorphic_binary_tree nat)) : bool := *) (* Definition eqb_optional_binary_tree_of_optional_nats (t1 t2 : option (polymorphic_binary_tree (option nat))) : bool := *) (* Definition eqb_binary_tree_of_optional_binary_tree_of_nats (t1 t2 : polymorphic_binary_tree (option (polymorphic_binary_tree nat))) : bool := *) (* Definition eqb_binary_tree_of_optional_binary_tree_of_optional_nats (t1 t2 : polymorphic_binary_tree (option (polymorphic_binary_tree (option nat)))) : bool := *) (* ********** *) Inductive option' (V : Type) : Type := | Some' : V -> option' V | None' : option' V. Print option'. Set Implicit Arguments. Inductive option'' (V : Type) : Type := | Some'' : V -> option'' V | None'' : option'' V. Print option''. Check (Some 10, Some true). Check (Some' nat 10, Some' bool true). Check (Some'' 10, Some'' true). Check None. Check None'. Check None''. Check (@None nat). Check (@None' nat). Check (@None'' nat). (* ********** *) (* end of week-01_polymorphism.v *)