Course7: 1- The Module System & 2- Understanding how the tactics for inductive types work.
A tiny introduction to Coq's Module system.
Module Type Moduletest.
(* ... *)
End Moduletest.
(* ... *)
End Moduletest.
The most basic use of Module is to prevent name cashes:
Module Mynat.
Inductive nat:=
|O: nat
|S: nat → nat.
End Mynat.
Module Mynat2.
Inductive nat:=
|O: nat
|S: nat → nat.
End Mynat2.
Print nat.
Print Mynat.nat.
Print Mynat2.nat.
But modules can be used in much richer ways
in order to structure data and logical theories.
Require Import List.
Import ListNotations.
Example of "Module Type", also known as a signature, or
an interface.
Module Type MYLIST.
Parameter t : Type → Type.
Parameter empty : ∀ A, t A.
Parameter cons : ∀ A, A → t A → t A.
Parameter decons : ∀ A, t A → option (A × t A).
Parameter length : ∀ A, t A → nat.
End MYLIST.
Example of an implementation of this signature.
With the "<:" syntax below, Coq checks that the definitions
in module MyList0 are compatibles with MYLIST.
But MyList0 is left unrestricted afterwards.
Module MyList0 <: MYLIST.
Definition t := list.
Definition empty {A} : list A := [].
Definition cons := @List.cons.
Definition decons {A} (l:list A) :=
match l with
| [] ⇒ None
| x::l ⇒ Some (x,l)
end.
Definition length := @List.length.
(* Note: the implementation can contain extra stuff not mentionned in
the signature. *)
Definition truc := 1 + 2.
End MyList0.
Print MyList0.t.
Check MyList0.truc.
Compute (MyList0.cons _ 1 MyList0.empty).
The problem is that the interface specification
allows other implementations which do not correspond to
our idea:
Module MyList1 <: MYLIST.
Definition t := list.
Definition empty {A} : list A := [].
Definition cons {A} (a: A) (l: list A) := l.
Definition decons {A} (l:list A) : option (A × t A):= None.
Definition length {A} (l: list A) := 42.
End MyList1.
In fact, a signature can contain contain declaration
of computational data as well as logical data as well as
the declaration as a proper name space.
Module Type Moduletest'.
Parameter A: Set.
Parameter inf sup: nat.
Axiom Bounds : inf ≤ sup.
End Moduletest'.
We can use this to make the specification more precise,
in the module signature:
Module Type MYLOGICALLIST.
Parameter t : Type → Type.
Parameter empty : ∀ A, t A.
Parameter cons : ∀ A, A → t A → t A.
Parameter decons : ∀ A, t A → option (A × t A).
Parameter length : ∀ A, t A → nat.
Axiom empty_def: ∀ A, decons A (empty A) = None.
Axiom cons_decons: ∀ A, ∀ (a: A), ∀ (l: t A),
decons A (cons A a l) = Some (a, l).
Axiom length_cons : ∀ A, ∀ (a: A), ∀ (l: t A),
length A (cons A a l) = S(length A l).
End MYLOGICALLIST.
Module MyList <: MYLOGICALLIST.
(* data / computational component *)
Definition t := list.
Definition empty {A} : list A := [].
Definition cons := @List.cons.
Definition decons {A} (l:list A) :=
match l with
| [] ⇒ None
| x::l ⇒ Some (x,l)
end.
Definition length := @List.length.
(* logical component *)
Theorem empty_def: ∀ A: Type, @decons A (@empty A) = None.
Proof. trivial. Qed.
Theorem cons_decons: ∀ A, ∀ (a: A), ∀ (l: t A),
decons (cons A a l) = Some (a, l).
Proof. trivial. Qed.
Theorem length_cons : ∀ A, ∀ (a: A), ∀ (l: t A),
@length A (@cons A a l) = S(@length A l).
Proof. trivial. Qed.
End MyList.
Print MyList.t.
Compute (MyList.cons _ 1 MyList.empty).
Print MyList.length_cons.
Now, if we use the syntax ":" below instead of "<:" in the definition
of a module, all internal details will be masked afterwards, hidden
by MYLOGICALLIST, and only the information given by MYLOGICALLIST
will be available.
In particular, this will prevent here any computation, since
the body of the definitions will be inaccessible.
So in Coq this ":" is usually far too restrictive, unlike in
languages like OCaml.
Module RestrictedMyList : MYLOGICALLIST := MyList.
Print RestrictedMyList.t.
Compute (RestrictedMyList.cons _ 1 (RestrictedMyList.empty _)).
Fail Check RestrictedMyList.truc.
A "functor" ("foncteur" in French) is a module parametrized
by another module.
Example: starting from MYLOGICALLIST, one may propose some head and
tail functions.
Module HeadTail (M:MYLOGICALLIST).
Definition head {A} : M.t A → option A :=
fun l ⇒ match M.decons _ l with None ⇒ None | Some (x,l') ⇒ Some x end.
Definition tail {A} : M.t A → option (M.t A) :=
fun l ⇒ match M.decons _ l with None ⇒ None | Some (x,l') ⇒ Some l' end.
End HeadTail.
For now, this does not create any new concrete functions.
Fail Print HeadTail.head.
But we can use this in a generic way on any implementation of MYLOGICALLIST.
Module MyListHdTl := HeadTail(MyList).
Print MyListHdTl.head.
Compute MyListHdTl.head [1;2].
(* But we do not have anymore access to the components from MyList... *)
Fail Print MyListHdTl.cons.
We can even extend a module, via a notion of inclusion.
Module MyList2.
Include MyList.
Include HeadTail(MyList).
End MyList2.
Print MyList2.head.
Print MyList2.cons.
Lighter syntax for the same thing:
Module MyList3 := MyList <+ HeadTail.
Another example of functor: starting from a first module
satisfying interface Foo, we could build another one for which
the length function is working in constant time.
For that we store somewhere this size, and update it after
all operations. That's a typical example of time vs. space tradeoff.
Module FastLength (M:MYLIST) <: MYLIST.
Definition t A := (M.t A × nat)%type.
Definition empty A := (M.empty A, 0).
Definition cons A x (l:t A) :=
let (l,n) := l in
(M.cons A x l, S n).
Definition decons A (l:t A) :=
let (l,n) := l in
match M.decons A l with
| None ⇒ None
| Some (x,l) ⇒ Some (x,(l,pred n))
end.
Definition length A (l:t A) := snd l.
End FastLength.
Tactics for inductive types:
- case : performs case analyis on an inductive hypothesis, without induction;
- destruct : a more elaborate tactic to do case analysis;
- elim : perform inductive reasoning;
- induction tactic: perform an improved inductive reasoning;
- simpl : applies computation rules;
- rewrite H : rewrite a term t with a term u when H has type t=u, variant: rewrite <- H;
- injection : all inductive constructors are injective, from H : S x = S y then injection H provides x = y;
- discriminate : all inductive constructors are orthogonal, from H : S x = O then discriminate H or just discriminate proves False (hence anything);
- f_equal : proves a goal f x = f y, as long as sub-goal x = y could then be proved. Sort of dual to injection, except that it works for any function f, not only for inductive constructors.
More on tactics on inductive types.
Case analysis
Section CaseAnalysis.
Parameter A B C D : Prop.
Lemma Case1: (A → B ∨ C) → D.
Proof.
intro H. case H.
Show 1.
Show 2.
Show 3.
Restart.
intro H. destruct H.
Show 1.
Show 2.
Show 3.
Restart.
intro H. destruct H as [H1 | H2].
Show 3.
Abort.
End CaseAnalysis.
Print nat_ind.
elim
induction n
Require Import Arith.
Theorem le_plus_minus' : ∀ n m:nat, m ≤ n → n = m+(n-m).
Proof.
induction n.
Show 2.
+ intros m Hm.
pose Nat.le_0_r as H'.
destruct (H' m) as [H0 H1]. rewrite H0.
- reflexivity.
- assumption.
+ intros m Hm.
SearchPattern (_ ≤ S _).
destruct m as [| n'].
- simpl. reflexivity.
- simpl. f_equal. apply (IHn n').
SearchPattern (S ?X ≤ S ?Y → ?X ≤ ?Y).
apply le_S_n. assumption.
Qed.
Require Import List.
Import ListNotations.
Lemma orthognal_S_O_discr : ∀ n, S n = 0 → False.
Proof. intros. discriminate. Qed.
Definition discr (m : nat) : Prop :=
match m with
| O ⇒ True
| S _ ⇒ False
end.
Compute discr 0. (* True *)
Compute discr 1. (* False *)
Lemma orthognal_S_O : ∀ n, S n = 0 → False.
(* said otherwise : forall n, S n <> 0. *)
(* reminder : <> is the negation of = *)
(* ~ A is A -> False *)
Proof.
intros.
discriminate.
Undo.
change (discr (S n)). (* converts the statement into something convertible*)
rewrite H.
simpl.
constructor. (* or: exact I *)
Show Proof.
Qed.
for list
Definition discr_cons_nil {A}(l:list A) :=
match l with
| nil ⇒ True
| cons _ _ ⇒ False
end.
Lemma orthognal_cons_nil {A}: ∀ (x:A) l, cons x l = nil → False.
Proof.
discriminate.
Undo.
intros.
change (discr_cons_nil (cons x l)).
rewrite H. constructor.
Qed.
(* for nat : *)
Lemma inject_S : ∀ n m, S n = S m → n = m.
Proof.
intros.
injection H.
trivial.
Qed.
(* We need a projection of the first argument of S:
That's the predecessor Nat.pred ! *)
Definition exhibit_succ (n:nat) : nat :=
match n with
| S x ⇒ x
| O ⇒ O
end.
Compute exhibit_succ 5. (* 4 *)
Lemma inject_S : ∀ n m, S n = S m → n = m.
Proof.
intros.
(* injection H. *)
change (exhibit_succ (S n) = exhibit_succ (S m)).
rewrite H.
reflexivity.
Show Proof.
Qed.
(* Same, for lists.
First, a projection of the first argument of cons
(We need a way to fill the other constructor, here nil,
with something of the right type, here (default:A)
*)
Definition proj1_cons {A} (l:list A) (default:A) :=
match l with
| cons x l ⇒ x
| nil ⇒ default
end.
Lemma inject_cons {A}: ∀ (x x':A) l l',
x::l = x'::l' → x=x'.
Proof.
intros.
change (proj1_cons (x::l) x = proj1_cons (x'::l') x).
rewrite H.
reflexivity.
Qed.
(* Second, a projection of the second argument.
Easier, no need for "default" this time. *)
Definition proj2_cons {A} (l:list A) :=
match l with
| cons x l ⇒ l
| nil ⇒ nil
end.
Lemma inject_cons' {A}: ∀ (x x':A) l l',
x::l = x'::l' → l=l'.
Proof.
intros.
change (proj2_cons (x::l) = proj2_cons (x'::l')).
rewrite H.
reflexivity.
Qed.
This page has been generated by coqdoc