Library tac


Ltac rdup :=
  repeat match goal with [ H1 : ?P, H2 : ?P |- _ ] => clear H2 end.

Ltac des_and H x K :=
  let al := fresh x in
    let ar := fresh in
      destruct H as [al ar];
        try K al;
          try K ar
.

Ltac des' H :=
  let t := type of H in
    match t with
      | _ /\ _ => let al := fresh "al" in des_and H al des'
      | (_ * _)%type => let al := fresh "al" in des_and H al des'
      | exists x, _ => des_and H x des'
      | { x | _ } => des_and H x des'
      | { x : _ & _} => des_and H x des'
    end.

Ltac des H := des' H; rdup.

Ltac desi trm :=
  let Hdesi := fresh "Hdesi" in
    generalize trm;
      intro Hdesi;
        des Hdesi
.

Ltac desa := repeat match goal with H : _ |- _ => des H end.

Ltac hintro :=
  let H := fresh in
    intro H;
      des H.

Ltac hintros := repeat hintro.

Lemma _use : forall P Q R : Type,
  P -> (P -> Q -> R) -> ((P -> Q) -> R).
Proof.
  tauto.
Qed.

Ltac use' H OK KO :=
  let t := type of H in
    match t with
      | ?P -> ?Q =>
          refine (_use P Q _ _ _ H) ; [ |
            
              let HQ := fresh "HQ" in
                intros _ HQ; try OK HQ
          ]
      | _ => KO H
    end.

Ltac use H :=
  let OK H' := (use H'; clear H') in
    let KO H' := des H' in
      use' H OK KO;
      rdup.

Lemma _euse : forall (A: Type) (B: A -> Type) (C: Type)
  a, (B a -> C) -> ((forall a', B a') -> C).
Proof.
  intros.
  apply X.
  apply X0.
Qed.

Ltac euse' H :=
  let t := type of H in
    match t with
      | _ -> _ =>
        let OK H' := first [ (euse' H'; clear H') | des H' ] in
          let KO H' := des H' in
            use' H OK KO
      | forall a : _, _ =>
        let H' := fresh "Heuse" in
          refine (_euse _ _ _ _ _ H);
            intro H';
              euse' H';
              clear H'
    end.

Ltac euse H := first [ euse' H | des H ]; rdup.