Library heap_tactic

Require Import List.
Require Import ZArith.
Require Import List.
Require Import EqNat.
Require Import Classical.

Require Import util.
Require Import heap.

Open Local Scope heap_scope.

Lemma disjoint_up: forall x x1 x2 x3,
  x = x1 +++ x3 -> x1 # x3 -> x # x2 -> x1 # x2.

Lemma disjoint_up': forall x x1 x2 x3,
  x = heap.union x1 x3 -> heap.disjoint x1 x3 -> heap.disjoint x x2 -> heap.disjoint x3 x2.

Ltac Disjoint_heap :=
  match goal with
    | id: (?x1 +++ ?x2) # ?x |- _ => generalize (heap.disjoint_union' x1 x2 x (heap.disjoint_com (x1 +++ x2) x id)); intro Disjoint_heapA; inversion_clear Disjoint_heapA; clear id; Disjoint_heap
    | id: ?x # (?x1 +++ ?x2) |- _ => generalize (heap.disjoint_union' x1 x2 x id); intro Disjoint_heapA; inversion_clear Disjoint_heapA; clear id; Disjoint_heap
    | |- (?x1 +++ ?x2) # ?x3 => eapply heap.disjoint_com; eapply heap.disjoint_union; split; [ (Disjoint_simpl || Disjoint_heap) | (Disjoint_simpl || Disjoint_heap)]
    | |- ?x3 # (?x1 +++ ?x2) => eapply heap.disjoint_union; split; [ (Disjoint_simpl || Disjoint_heap) | (Disjoint_simpl || Disjoint_heap)]
    | |- ?x1 # ?x2 => Disjoint_up
  end
with
  Disjoint_up :=
  (
    Disjoint_simpl || (Disjoint_up_left) || (Disjoint_up_right)
  )
with
  Disjoint_up_left :=
  match goal with
    | id1: ?X1 = (?x1 +++ ?x1') |- ?x1 # ?x2 => apply (disjoint_up X1 x1 x2 x1' id1); [(Disjoint_simpl || Disjoint_heap)|(Disjoint_simpl || Disjoint_heap)]
    | id1: ?X1 = (?x1 +++ ?x1') |- ?x1' # ?x2 => apply (disjoint_up' X1 x1 x2 x1' id1) ; [(Disjoint_simpl || Disjoint_heap)|(Disjoint_simpl || Disjoint_heap)]
    | |- ?x1 # ?x2 => (Disjoint_simpl)
  end
with
  Disjoint_up_right :=
  match goal with
    | id1: ?X1 = (?x2 +++ ?x2') |- ?x1 # ?x2 => apply heap.disjoint_com; apply (disjoint_up X1 x2 x1 x2' id1); [(Disjoint_simpl || Disjoint_heap)|(Disjoint_simpl || Disjoint_heap)]
    | id1: ?X1 = (?x2 +++ ?x2') |- ?x1 # ?x2' => apply heap.disjoint_com; apply (disjoint_up' X1 x2 x1 x2' id1) ; [(Disjoint_simpl || Disjoint_heap)|(Disjoint_simpl || Disjoint_heap)]
    | |- ?x1 # ?x2 => (Disjoint_simpl)
  end
with
  Disjoint_simpl :=
  match goal with
    | id : ?x1 # ?x2 |- ?x1 # ?x2 => auto
    | id : ?x2 # ?x1 |- ?x1 # ?x2 => apply heap.disjoint_com; auto
    | |- ?x1 # heap.emp => apply (heap.disjoint_emp x1)
    | |- heap.emp # ?x1 => apply heap.disjoint_com; apply (heap.disjoint_emp x1)
    | |- heap.singleton ?l1 ?v1 # heap.singleton ?l2 ?v2 => eapply heap.disjoint_singleton; omega
  end.

Lemma tactic_test_disjoint: forall h h1 h2 h11 h12 h21 h22 h111 h112 h121 h122 h211 h212 h221 h222,
  h1 # h2 ->
  h11 # h12 ->
  h21 # h22 ->
  h111 # h112 ->
  h121 # h122 ->
  h211 # h212 ->
  h221 # h222 ->
  h = (heap.union h1 h2) ->
  h1 = (heap.union h11 h12) ->
  h2 = (heap.union h21 h22) ->
  h11 = (heap.union h111 h112) ->
  h12 = (heap.union h121 h122) ->
  h21 = (heap.union h211 h212) ->
  h22 = (heap.union h221 h222) ->
  (h112 +++ h221) # h222.

Lemma tactic_test_disjoint2: forall h h1 h2,
  heap.disjoint h (heap.union h1 h2) ->
  heap.disjoint h h2 .

Lemma equal_tactic_l1: forall h1 h2 h3 h4,
  heap.disjoint h1 h2 ->
  (heap.union h2 h1) = (heap.union h3 h4) ->
  (heap.union h1 h2) = (heap.union h3 h4).

Lemma equal_tactic_l1': forall h1 h2 h3 h4,
  heap.disjoint h3 h4 ->
  (heap.union h1 h2) = (heap.union h4 h3) ->
  (heap.union h1 h2) = (heap.union h3 h4).

Lemma equal_tactic_l2: forall h1 h2 h3 H,
  H = heap.union (heap.union h1 h2) h3 ->
  H = heap.union h1 (heap.union h2 h3).

Lemma equal_tactic_l2': forall h1 h2 h3,
  heap.union (heap.union h1 h2) h3 = heap.union h1 (heap.union h2 h3).

Lemma equal_tactic_l3: forall h1 h2 h3 H,
        (heap.union (heap.union h1 h2) h3) = H ->
              (heap.union h1 (heap.union h2 h3)) = H.

Lemma equal_tactic_l4: forall x1 x2 h1 h2 H,
        H = (heap.union x1 x2) ->
              (heap.union x1 x2) = (heap.union h1 h2) ->
                  (heap.union h1 h2) = H.

Lemma equal_tactic_l4': forall x1 h1 h2 H,
        H = x1 ->
              x1 = (heap.union h1 h2) ->
                  (heap.union h1 h2) = H.

Lemma equal_tactic_l5: forall x1 x2 h1 h2 H,
        H = (heap.union x1 x2) ->
              (heap.union x1 x2) = (heap.union h1 h2) ->
                  H = (heap.union h1 h2).

Lemma equal_tactic_l5': forall x1 h1 h2 H,
        H = x1 ->
              x1 = (heap.union h1 h2) ->
                  H = (heap.union h1 h2).

Lemma equal_tactic_l6: forall X Y,
     X = Y ->
         (heap.union X heap.emp) = Y.

Lemma equal_tactic_l7: forall X Y,
      Y = X ->
         (heap.union heap.emp X ) = Y.

Lemma equal_tactic_l8: forall x1 x2 h1 X Y,
  X = x1 +++ x2 ->
  h1 # X ->
  x1 # x2 ->
  h1 +++ (x1 +++ x2) = Y ->
  h1 +++ X = Y.

Lemma equal_tactic_l8': forall x1 h1 X Y,
  X = x1 ->
  heap.disjoint h1 X ->
  (heap.union h1 x1) = Y ->
  (heap.union h1 X) = Y.

Lemma equal_tactic_l9: forall X Y h1,
            h1 = X ->
               X = Y ->
                  (heap.union X h1) = (heap.union Y h1).

Lemma equal_tactic_l9': forall X Y x1 x2,
      x1 = x2 ->
            heap.disjoint x1 X ->
            heap.disjoint x2 X ->
               X = Y ->
                  (heap.union X x1) = (heap.union Y x2).

Lemma equal_tactic_l10: forall X h1 h2,
            h1 = h2 ->
                  (heap.union X h1) = (heap.union X h2).

Ltac Equal_heap :=
     match goal with
        | |- (?h1 +++ ?h2) = (?h1 +++ ?h3) => apply (equal_tactic_l10 h1 h2 h3); Equal_heap
        | |- ?h1 = ?h1 => auto
        | |- (heap.emp +++ ?h1) = ?h2 => apply (equal_tactic_l7 h1 h2); [Equal_heap]
        | |- (?h1 +++ heap.emp) = ?h2 => apply (equal_tactic_l6 h1 h2); [Equal_heap]
        | |- ?h2 = (heap.emp +++ ?h1) => symmetry; apply (equal_tactic_l7 h1 h2); [Equal_heap]
        | |- ?h2 = (?h1 +++ heap.emp) => symmetry;apply (equal_tactic_l6 h1 h2); [Equal_heap]
        | id: ?X = (?x1 +++ ?x2) |- (?X +++ ?x1') = (?Y +++ ?x2') => rewrite id ; Equal_heap
        | id: ?Y = (?y1 +++ ?y2) |- (?X +++ ?x1') = (?Y +++ ?x2') => rewrite id ; Equal_heap
        | id: ?X = (?x1 +++ ?x2) |- ?X = (?Y +++ ?x2') => rewrite id ; Equal_heap
        | id: ?Y = (?y1 +++ ?y2) |- (?X +++ ?x1') = ?Y => rewrite id ; Equal_heap
        | |- ((?h1 +++ ?h2) +++ ?h3) = ?X => rewrite (equal_tactic_l2' h1 h2 h3); [Equal_heap]
        | |- ?X = ((?h1 +++ ?h2) +++ ?h3) => rewrite (equal_tactic_l2' h1 h2 h3); [Equal_heap]
        | |- (?h1 +++ ?h2) = (?h3 +++ ?h4) => apply (equal_tactic_l1 h1 h2 h3 h4); [Disjoint_heap | Equal_heap]
        | id1: ?h1 = ?h2 |- ?h1 = ?h3 => rewrite id1; Equal_heap
        | id1: ?h1 = ?h2 |- ?h3 = ?h1 => rewrite id1; Equal_heap
     end.