Lvc.Infra.AutoIndTac

Require Import List.
Require Export Smpl LvcPlugin LengthEq MoreInversion.

Set Implicit Arguments.


Ltac revert_except E :=
  repeat match goal with
           [H : _ |- _] ⇒
           match H with
             | Efail 1
             | _revert H
               end
         end.

Ltac clear_except E :=
  repeat match goal with
           [H : _ |- _] ⇒
           match H with
             | Efail
             | _clear H
               end
         end.

Ltac clear_all :=
  repeat match goal with
           [H : _ |- _] ⇒ clear H
         end.

Ltac revert_all :=
  repeat match goal with
           [H : _ |- _] ⇒ revert H
         end.


Class DoNotRemember (T:Type) := DNR { Q:Type }.

Ltac remember_arguments E :=
  let tac t x := ( try (is_param t 0; fail 1);
                  try (is_var x; fail 1);
                  try (assert (DoNotRemember x) by eauto with typeclass_instances; fail 1 );
                  remember (x))
  in
  repeat (match type of E with
    | ?t ?x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _tac t x
    | ?t ?x _ _ _ _ _ _ _ _ _ _ _ _ _ _tac t x
    | ?t ?x _ _ _ _ _ _ _ _ _ _ _ _ _tac t x
    | ?t ?x _ _ _ _ _ _ _ _ _ _ _ _tac t x
    | ?t ?x _ _ _ _ _ _ _ _ _ _ _tac t x
    | ?t ?x _ _ _ _ _ _ _ _ _ _tac t x
    | ?t ?x _ _ _ _ _ _ _ _ _tac t x
    | ?t ?x _ _ _ _ _ _ _ _tac t x
    | ?t ?x _ _ _ _ _ _ _tac t x
    | ?t ?x _ _ _ _ _ _tac t x
    | ?t ?x _ _ _ _ _tac t x
    | ?t ?x _ _ _ _tac t x
    | ?t ?x _ _ _tac t x
    | ?t ?x _ _tac t x
    | ?t ?x _tac t x
    | ?t ?xtac t x
  end).

Ltac clear_dup :=

  match goal with
  | [ H : ?X |- _ ] ⇒
    match goal with
    | [ H' : ?Y |- _ ] ⇒
      match H with
      | H'fail 2
      | _unify X Y ; (clear H' || clear H)
      end
    end
  end.

Ltac clear_if_dup H :=
  match type of H with
    | ?X
      match goal with
      | [ H' : ?Y |- _ ] ⇒
        match H with
        | H'fail 2
        | _unify X Y ; (clear H' || clear H)
        end
      end
  end.

Ltac inductify H :=
  match type of H with
  | length _ = length _eapply length_length_eq in H
  end.

Ltac expose_inductive H :=
  first [ is_inductive H | hnf in H ].

Tactic Notation "general" "induction" hyp(H) :=
  (try inductify H);
  expose_inductive H;
  remember_arguments H; revert_except H;
  induction H; intros; (try clear_trivial_eqs);
  repeat (smpl inversion_cleanup).

Tactic Notation "indros" :=
  intros; (try inv_eqs); (try clear_trivial_eqs).

Module Test.

  Inductive decreasing : list nat Prop :=
  | base : decreasing nil
  | step m n L : decreasing (n::L) n m decreasing (m :: n :: L).

  Lemma all_zero_by_hand L
    : decreasing (0::L) x, In x L x = 0.
  Proof.
    intros. remember (0::L).
    revert dependent L. revert x. induction H; intros.
    inversion Heql.
    inversion Heql. subst. inversion H0; subst; firstorder.
  Qed.

  Lemma all_zero L
    : decreasing (0::L) x, In x L x = 0.
  Proof.
    intros. general induction H.
    firstorder.
  Qed.

End Test.