comparison src/Equality.v @ 123:9ccd215e5112

ext_eq
author Adam Chlipala <adamc@hcoop.net>
date Sat, 18 Oct 2008 17:36:08 -0400
parents 7cbf0376702f
children 57eaceb085f6
comparison
equal deleted inserted replaced
122:7cbf0376702f 123:9ccd215e5112
8 *) 8 *)
9 9
10 (* begin hide *) 10 (* begin hide *)
11 Require Import Eqdep JMeq List. 11 Require Import Eqdep JMeq List.
12 12
13 Require Import Tactics. 13 Require Import MoreSpecif Tactics.
14 14
15 Set Implicit Arguments. 15 Set Implicit Arguments.
16 (* end hide *) 16 (* end hide *)
17 17
18 18
721 ]] *) 721 ]] *)
722 722
723 rewrite (UIP_refl _ _ x0); reflexivity. 723 rewrite (UIP_refl _ _ x0); reflexivity.
724 Qed. 724 Qed.
725 725
726 (** We see that, in a very formal sense, we are free to switch back and forth between the two styles of proofs about equality proofs. One style may be more convenient than the other for some proofs, but we can always intercovert between our results. The style that does not use heterogeneous equality may be preferable in cases where many results do not require the tricks of this chapter, since then the use of axioms is avoided altogether for the simple cases, and a wider audience will be able to follow those "simple" proofs. On the other hand, heterogeneous equality often makes for shorter and more readable theorem statements. *) 726 (** We see that, in a very formal sense, we are free to switch back and forth between the two styles of proofs about equality proofs. One style may be more convenient than the other for some proofs, but we can always intercovert between our results. The style that does not use heterogeneous equality may be preferable in cases where many results do not require the tricks of this chapter, since then the use of axioms is avoided altogether for the simple cases, and a wider audience will be able to follow those "simple" proofs. On the other hand, heterogeneous equality often makes for shorter and more readable theorem statements.
727
728 It is worth remarking that it is possible to avoid axioms altogether for equalities on types with decidable equality. The [Eqdep_dec] module of the standard library contains a parametric proof of [UIP_refl] for such cases. *)
729
730
731 (** * Equality of Functions *)
732
733 (** The following seems like a reasonable theorem to want to hold, and it does hold in set theory. [[
734
735 Theorem S_eta : S = (fun n => S n).
736
737 Unfortunately, this theorem is not provable in CIC without additional axioms. None of the definitional equality rules force function equality to be %\textit{%#<i>#extensional#</i>#%}%. That is, the fact that two functions return equal results on equal inputs does not imply that the functions are equal. We %\textit{%#<i>#can#</i>#%}% assert function extensionality as an axiom. *)
738
739 Axiom ext_eq : forall A B (f g : A -> B),
740 (forall x, f x = g x)
741 -> f = g.
742
743 (** This axiom has been verified metatheoretically to be consistent with CIC and the two equality axioms we considered previously. With it, the proof of [S_eta] is trivial. *)
744
745 Theorem S_eta : S = (fun n => S n).
746 apply ext_eq; reflexivity.
747 Qed.
748
749 (** The same axiom can help us prove equality of types, where we need to "reason under quantifiers." *)
750
751 Theorem forall_eq : (forall x : nat, match x with
752 | O => True
753 | S _ => True
754 end)
755 = (forall _ : nat, True).
756
757 (** There are no immediate opportunities to apply [ext_eq], but we can use [change] to fix that. *)
758
759 change ((forall x : nat, (fun x => match x with
760 | 0 => True
761 | S _ => True
762 end) x) = (nat -> True)).
763 rewrite (ext_eq (fun x => match x with
764 | 0 => True
765 | S _ => True
766 end) (fun _ => True)).
767 (** [[
768
769 2 subgoals
770
771 ============================
772 (nat -> True) = (nat -> True)
773
774 subgoal 2 is:
775 forall x : nat, match x with
776 | 0 => True
777 | S _ => True
778 end = True
779 ]] *)
780
781
782 reflexivity.
783
784 destruct x; constructor.
785 Qed.
786