1 2 3 4 5 6 7 8 9 10 11 12 13 14 15 16 17 18 19 20 21 22 23 24 25 26 27 28 29 30 31 32 33 34 35 36 37 38 39 40 41 42 43 44 45 46 47 48 49 50 51 52 53 54 55 56 57 58 59 60 61 62 63 64 65 66 67
|
Module Type Sub.
Axiom Refl1 : forall x : nat, x = x.
Axiom Refl2 : forall x : nat, x = x.
Axiom Refl3 : forall x : nat, x = x.
Inductive T : Set :=
A : T.
End Sub.
Module Type Main.
Declare Module M: Sub.
End Main.
Module A <: Main.
Module M <: Sub.
Lemma Refl1 : forall x : nat, x = x.
intros; reflexivity.
Qed.
Axiom Refl2 : forall x : nat, x = x.
Lemma Refl3 : forall x : nat, x = x.
intros; reflexivity.
Defined.
Inductive T : Set :=
A : T.
End M.
End A.
(* first test *)
Module F (S: Sub).
Module M := S.
End F.
Module B <: Main with Module M:=A.M := F A.M.
(* second test *)
Lemma r1 : (A.M.Refl1 = B.M.Refl1).
Proof.
reflexivity.
Qed.
Lemma r2 : (A.M.Refl2 = B.M.Refl2).
Proof.
reflexivity.
Qed.
Lemma r3 : (A.M.Refl3 = B.M.Refl3).
Proof.
reflexivity.
Qed.
Lemma t : (A.M.T = B.M.T).
Proof.
reflexivity.
Qed.
Lemma a : (A.M.A = B.M.A).
Proof.
reflexivity.
Qed.
|