]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/coq-contribs/LAMBDA-TYPES/base_hints.v
contribution about \lambda-\delta
[helm.git] / helm / coq-contribs / LAMBDA-TYPES / base_hints.v
diff --git a/helm/coq-contribs/LAMBDA-TYPES/base_hints.v b/helm/coq-contribs/LAMBDA-TYPES/base_hints.v
new file mode 100644 (file)
index 0000000..a364406
--- /dev/null
@@ -0,0 +1,262 @@
+(*#* #stop file *)
+
+Require Arith.
+Require base_tactics.
+
+(* eq ***********************************************************************)
+
+Hint eq : ltlc := Constructors eq.
+
+Hint f1N : ltlc := Resolve (f_equal nat).
+
+Hint f2NN : ltlc := Resolve (f_equal2 nat nat).
+
+Hints Resolve sym_equal : ltlc.
+
+Hints Resolve plus_sym plus_n_Sm plus_assoc_r simpl_plus_l : ltlc.
+
+Hints Resolve minus_n_O : ltlc.
+
+(* le ***********************************************************************)
+
+Hint le : ltlc := Constructors le.
+
+Hints Resolve le_O_n le_n_S le_S_n le_trans : ltlc.
+
+Hints Resolve lt_le_S : ltlc.
+
+Hints Resolve le_plus_plus le_plus_trans le_plus_l le_plus_r : ltlc.
+
+(* lt ***********************************************************************)
+
+Hints Resolve lt_trans : ltlc.
+
+Hints Resolve lt_le_trans le_lt_n_Sm : ltlc.
+
+Hints Resolve lt_reg_r lt_le_plus_plus le_lt_plus_plus : ltlc.
+
+(* not **********************************************************************)
+
+Hints Resolve sym_not_equal : ltlc.
+
+(* missing in the standard library ******************************************)
+
+      Theorem simpl_plus_r: (n,m,p:?) (plus m n) = (plus p n) -> m = p.
+      Intros.
+      Apply (simpl_plus_l n).
+      Rewrite plus_sym.
+      Rewrite H; XAuto.
+      Qed.
+
+      Theorem minus_plus_r: (m,n:?) (minus (plus m n) n) = m.
+      Intros.
+      Rewrite plus_sym.
+      Apply minus_plus.
+      Qed.
+
+      Theorem plus_permute_2_in_3: (x,y,z:?) (plus (plus x y) z) = (plus (plus x z) y).
+      Intros.
+      Rewrite plus_assoc_r.
+      Rewrite (plus_sym y z).
+      Rewrite <- plus_assoc_r; XAuto.
+      Qed.
+
+      Theorem plus_permute_2_in_3_assoc: (n,h,k:?) (plus (plus n h) k) = (plus n (plus k h)).
+      Intros.
+      Rewrite plus_permute_2_in_3; Rewrite plus_assoc_l; XAuto.
+      Qed.
+
+      Theorem plus_O: (x,y:?) (plus x y) = (0) -> x = (O) /\ y = (O).
+      XElim x; [ XAuto | Intros; Inversion H0 ].
+      Qed.
+
+      Theorem minus_Sx_SO: (x:?) (minus (S x) (1)) = x.
+      Intros; Simpl; Rewrite <- minus_n_O; XAuto.
+      Qed.
+
+      Theorem eq_nat_dec: (i,j:nat) ~i=j \/ i=j.
+      XElim i; XElim j; Intros; XAuto.
+      Elim (H n0); XAuto.
+      Qed.
+
+      Theorem neq_eq_e: (i,j:nat; P:Prop) (~i=j -> P) -> (i=j -> P) -> P.
+      Intros.
+      Pose (eq_nat_dec i j).
+      XElim o; XAuto.
+      Qed.
+
+      Theorem le_false: (m,n:?; P:Prop) (le m n) -> (le (S n) m) -> P.
+      XElim m.
+(* case 1 : m = 0 *)
+      Intros; Inversion H0.
+(* case 2 : m > 0 *)
+      XElim n0; Intros.
+(* case 2.1 : n = 0 *)
+      Inversion H0.
+(* case 2.2 : n > 0 *)
+      Simpl in H1.
+      Apply (H n0); XAuto.
+      Qed.
+
+      Theorem le_plus_minus_sym: (n,m:?) (le n m) -> m = (plus (minus m n) n).
+      Intros.
+      Rewrite plus_sym; Apply le_plus_minus; XAuto.
+      Qed.
+
+      Theorem le_minus_minus: (x,y:?) (le x y) -> (z:?) (le y z) ->
+                              (le (minus y x) (minus z x)).
+      Intros.
+      EApply simpl_le_plus_l.
+      Rewrite le_plus_minus_r; [ Idtac | XAuto ].
+      Rewrite le_plus_minus_r; XEAuto.
+      Qed.
+
+      Theorem le_minus_plus: (z,x:?) (le z x) -> (y:?)
+                             (minus (plus x y) z) = (plus (minus x z) y).
+      XElim z.
+(* case 1 : z = 0 *)
+      Intros x H; Inversion H; XAuto.
+(* case 2 : z > 0 *)
+      Intros z; XElim x; Intros.
+(* case 2.1 : x = 0 *)
+      Inversion H0.
+(* case 2.2 : x > 0 *)
+      Simpl; XAuto.
+      Qed.
+
+      Theorem le_minus: (x,z,y:?) (le (plus x y) z) -> (le x (minus z y)).
+      Intros.
+      Rewrite <- (minus_plus_r x y); XAuto.
+      Apply le_minus_minus; XAuto.
+      Qed.
+
+      Theorem le_trans_plus_r: (x,y,z:?) (le (plus x y) z) -> (le y z).
+      Intros.
+      EApply le_trans; [ EApply le_plus_r | Idtac ]; XEAuto.
+      Qed.
+
+      Theorem le_gen_S: (m,x:?) (le (S m) x) ->
+                        (EX n | x = (S n) & (le m n)).
+      Intros; Inversion H; XEAuto.
+      Qed.
+
+      Theorem lt_x_plus_x_Sy: (x,y:?) (lt x (plus x (S y))).
+      Intros; Rewrite plus_sym; Simpl; XAuto.
+      Qed.
+
+      Theorem simpl_lt_plus_r: (p,n,m:?) (lt (plus n p) (plus m p)) -> (lt n m).
+      Intros.
+      EApply simpl_lt_plus_l.
+      Rewrite plus_sym in H; Rewrite (plus_sym m p) in H; Apply H.
+      Qed.
+
+      Theorem minus_x_Sy: (x,y:?) (lt y x) ->
+                          (minus x y) = (S (minus x (S y))).
+      XElim x.
+(* case 1 : x = 0 *)
+      Intros; Inversion H.
+(* case 2 : x > 0 *)
+      XElim y; Intros; Simpl.
+(* case 2.1 : y = 0 *)
+      Rewrite <- minus_n_O; XAuto.
+(* case 2.2 : y > 0 *)
+      Cut (lt n0 n); XAuto.
+      Qed.
+
+      Theorem lt_plus_minus: (x,y:?) (lt x y) ->
+                             y = (S (plus x (minus y (S x)))).
+      Intros.
+      Apply (le_plus_minus (S x) y); XAuto.
+      Qed.
+
+      Theorem lt_plus_minus_r: (x,y:?) (lt x y) ->
+                               y = (S (plus (minus y (S x)) x)).
+      Intros.
+      Rewrite plus_sym; Apply lt_plus_minus; XAuto.
+      Qed.
+
+      Theorem minus_x_SO: (x:?) (lt (0) x) -> x = (S (minus x (1))).
+      Intros.
+      Rewrite <- minus_x_Sy; [ Rewrite <- minus_n_O; XEAuto | XEAuto ].
+      Qed.
+
+      Theorem lt_le_minus: (x,y:?) (lt x y) -> (le x (minus y (1))).
+      Intros; Apply le_minus; Rewrite plus_sym; Simpl; XAuto.
+      Qed.
+
+      Theorem lt_le_e: (n,d:?; P:Prop)
+                       ((lt n d) -> P) -> ((le d n) -> P) -> P.
+      Intros.
+      Cut (le d n) \/ (lt n d); [ Intros H1; XElim H1; XAuto | Apply le_or_lt ].
+      Qed.
+
+      Theorem lt_eq_e: (x,y:?; P:Prop) ((lt x y) -> P) ->
+                       (x = y -> P) -> (le x y) -> P.
+      Intros.
+      LApply (le_lt_or_eq x y); [ Clear H1; Intros H1 | XAuto ].
+      XElim H1; XAuto.
+      Qed.
+
+      Theorem lt_eq_gt_e: (x,y:?; P:Prop) ((lt x y) -> P) ->
+                          (x = y -> P) -> ((lt y x) -> P) -> P.
+      Intros.
+      Apply (lt_le_e x y); [ XAuto | Intros ].
+      Apply (lt_eq_e y x); XAuto.
+      Qed.
+
+      Theorem lt_gen_S': (x,n:?) (lt x (S n)) ->
+                         x = (0) \/ (EX m | x = (S m) & (lt m n)).
+      XElim x; XEAuto.
+      Qed.
+
+Hints Resolve le_lt_trans : ltlc.
+
+Hints Resolve simpl_plus_r minus_plus_r minus_x_Sy
+              plus_permute_2_in_3 plus_permute_2_in_3_assoc : ltlc.
+
+Hints Resolve le_minus_minus le_minus_plus le_minus le_trans_plus_r : ltlc.
+
+Hints Resolve lt_x_plus_x_Sy simpl_lt_plus_r lt_le_minus lt_plus_minus
+              lt_plus_minus_r : ltlc.
+
+      Theorem lt_neq: (x,y:?) (lt x y) -> ~x=y.
+      Unfold not; Intros; Rewrite H0 in H; Clear H0 x.
+      LApply (lt_n_n y); XAuto.
+      Qed.
+
+Hints Resolve lt_neq : ltlc.
+
+      Theorem arith0: (h2,d2,n:?) (le (plus d2 h2) n) ->
+                      (h1:?) (le (plus d2 h1) (minus (plus n h1) h2)).
+      Intros.
+      Rewrite <- (minus_plus h2 (plus d2 h1)).
+      Apply le_minus_minus; [ XAuto | Idtac ].
+      Rewrite plus_assoc_l; Rewrite (plus_sym h2 d2); XAuto.
+      Qed.
+
+Hints Resolve arith0 : ltlc.
+
+      Tactic Definition EqFalse :=
+         Match Context With
+            [ H: ~?1=?1 |- ? ] ->
+               LApply H; [ Clear H; Intros H; Inversion H | XAuto ].
+
+      Tactic Definition PlusO :=
+         Match Context With
+            | [ H: (plus ?0 ?1) = (0) |- ? ] ->
+               LApply (plus_O ?0 ?1); [ Clear H; Intros H | XAuto ];
+               XElim H; Intros.
+
+      Tactic Definition SymEqual :=
+         Match Context With
+            | [ H: ?1 = ?2 |- ? ] ->
+               Cut ?2 = ?1; [ Clear H; Intros H | Apply sym_equal; XAuto ].
+
+      Tactic Definition LeLtGen :=
+         Match Context With
+            | [ H: (le (S ?1) ?2) |- ? ] ->
+               LApply (le_gen_S ?1 ?2); [ Clear H; Intros H | XAuto ];
+               XElim H; Intros
+            | [ H: (lt ?1 (S ?2)) |- ? ] ->
+               LApply (lt_gen_S' ?1 ?2); [ Clear H; Intros H | XAuto ];
+               XElim H; [ Intros | Intros H; XElim H; Intros ].