Library map

Require Export FunctionalExtensionality List.
Section Map.

Variable A : Type.

Section Map_values.

Variable B : Type.

Set Implicit Arguments.
Unset Strict Implicit.

Definition Map : Type := (A -> option B).

Definition lookup_strong (M : Map) (a : A) : option B := M a.

Definition lookup (M : Map) a b := lookup_strong M a = Some b.

Lemma lookup_func : forall M a b1, lookup M a b1 -> forall b2, lookup M a b2 -> b1 = b2.
Proof.
  unfold lookup.
  congruence.
Qed.

Definition in_dom (M : Map) a := exists b, lookup M a b.

Lemma in_dom_descr: forall (M: Map) a, in_dom M a -> {b | lookup M a b}.
Proof.
  intros.
  case_eq (lookup_strong M a).
   unfold lookup; eauto.
  intros.
  unfold in_dom, lookup in *.
  exfalso. destruct H. congruence.
Qed.

Lemma in_dom_dec: forall (M: Map) a, {in_dom M a} + {~ in_dom M a}.
Proof.
  intros.
  case_eq (lookup_strong M a).
   left. esplit. eassumption.
  right. intro. destruct H0. unfold lookup in *. congruence.
Qed.

Lemma lookup_in_dom: forall M a b, lookup M a b -> in_dom M a.
Proof.
  unfold in_dom. eauto.
Qed.

Lemma lookup_strong_none_not_in_dom: forall M a,
  lookup_strong M a = None -> ~ in_dom M a.
Proof.
  intros.
  intro.
  destruct H0.
  unfold lookup.
  congruence.
Qed.

Lemma not_in_dom_lookup_strong_none: forall M a,
  (~ in_dom M a) -> lookup_strong M a = None.
Proof.
  intros.
  case_eq (lookup_strong M a).
   intros.
   destruct H.
   eauto using lookup_in_dom.
  trivial.
Qed.

Lemma in_dom_ext: forall m1 m2
  (Hext: forall f, lookup_strong m1 f = lookup_strong m2 f) f,
  (in_dom m1 f <-> in_dom m2 f).
Proof.
  intros.
  unfold in_dom, lookup.
  rewrite Hext.
  tauto.
Qed.

Definition empty_map : Map := fun _ => None.

Definition merge (M1 M2 : Map) (a: A) :=
  match M1 a with
    | Some b => Some b
    | _ => M2 a
  end.

Lemma merge_empty: forall M, merge empty_map M = M.
Proof.
  intros.
  apply functional_extensionality.
  reflexivity.
Qed.


Definition disjoint (M1 M2 : Map) := forall a, in_dom M1 a -> in_dom M2 a -> False.

Lemma disjoint_empty : forall M, disjoint empty_map M.
Proof.
  unfold disjoint, empty_map.
  destruct 1.
  discriminate.
Qed.

Lemma disjoint_sym: forall m1 m2, disjoint m1 m2 -> disjoint m2 m1.
Proof. unfold disjoint; firstorder. Qed.

Definition subseteq (M1 M2 : Map) := forall a b, lookup M1 a b -> lookup M2 a b.

Lemma mergelookup :
 forall (M1 M2 : Map) a b,
 lookup (merge M1 M2) a b -> lookup M1 a b \/ lookup M2 a b.
Proof.
  unfold lookup.
  unfold merge.
  intros until b.
  unfold lookup_strong.
  case_eq (M1 a); auto.
Qed.

Lemma mergelookup' :
 forall (M1 M2 : Map) a b,
 lookup (merge M1 M2) a b -> exists M, lookup M a b /\ (M = M1 \/ M = M2).
Proof.
  intros.
  destruct (mergelookup H); eauto.
Qed.

Lemma mergelookup_1:
 forall (M1 M2 : Map) a b,
 lookup (merge M1 M2) a b ->
 ~ in_dom M2 a ->
 lookup M1 a b.
Proof.
  intros.
  destruct (mergelookup H).
   assumption.
  destruct H0. esplit. eassumption.
Qed.

Lemma mergelookup_2:
 forall (M1 M2 : Map) a b,
 lookup (merge M1 M2) a b ->
 ~ in_dom M1 a ->
 lookup M2 a b.
Proof.
  intros.
  destruct (mergelookup H).
   destruct H0. esplit. eassumption.
  assumption.
Qed.

Corollary merge_indom : forall (M1 M2 : Map) a,
  in_dom (merge M1 M2) a -> in_dom M1 a \/ in_dom M2 a.
Proof.
  destruct 1.
  eapply mergelookup in H.
  unfold in_dom; destruct H; eauto.
Qed.

Lemma lookup_merge_lookup_1 :
  forall M1 M2 a b, lookup M1 a b -> lookup (merge M1 M2) a b.
Proof.
  unfold merge.
  unfold lookup.
  unfold lookup_strong.
  intros.
  rewrite H.
  trivial.
Qed.

Lemma lookup_merge_lookup_2 :
  forall M1 a, ~ in_dom M1 a -> forall M2 b, lookup M2 a b -> lookup (merge M1 M2) a b.
Proof.
  intros until 1.
  generalize (not_in_dom_lookup_strong_none H).
  unfold merge, lookup, lookup_strong.
  intro.
  rewrite H0.
  tauto.
Qed.

Corollary lookup_disj_merge_lookup_2 :
  forall M1 M2 a b, lookup M2 a b -> disjoint M1 M2 -> lookup (merge M1 M2) a b.
Proof.
  intros.
  eapply lookup_merge_lookup_2.
   intro.
   edestruct H0.
   eassumption.
   eapply lookup_in_dom.
   eassumption.
  assumption.
Qed.

Corollary lookup_disj_merge:
  forall M1 M2 M (HM: M = M1 \/ M = M2) a b, lookup M a b -> disjoint M1 M2 -> lookup (merge M1 M2) a b.
Proof.
  destruct 1; subst; eauto using lookup_merge_lookup_1, lookup_disj_merge_lookup_2.
Qed.

Lemma indom_merge_indom_1 :
  forall M1 M2 a, in_dom M1 a -> in_dom (merge M1 M2) a.
Proof.
  unfold merge, in_dom, lookup, lookup_strong.
  destruct 1.
  rewrite H.
  eauto.
Qed.

Lemma indom_merge_indom_2 :
  forall M1 M2 a, in_dom M2 a -> in_dom (merge M1 M2) a.
Proof.
  unfold merge, in_dom, lookup, lookup_strong.
  destruct 1.
  case_eq (M1 a); eauto.
Qed.

Corollary indom_merge:
  forall M1 M2 M (HM: M = M1 \/ M = M2) a, in_dom M a -> in_dom (merge M1 M2) a.
Proof.
  destruct 1; subst; eauto using indom_merge_indom_1, indom_merge_indom_2.
Qed.

Lemma indom_disj_notindom :
  forall M1 M2 a, in_dom M1 a -> disjoint M1 M2 -> ~ in_dom M2 a.
Proof.
  unfold disjoint; firstorder.
Qed.

Lemma disj_sym : forall M1 M2, disjoint M1 M2 -> disjoint M2 M1.
Proof.
  unfold disjoint; firstorder.
Qed.

Lemma disj_merge_sym :
  forall M1 M2, disjoint M1 M2 -> merge M1 M2 = merge M2 M1.
Proof.
intros; extensionality a; unfold merge, disjoint, in_dom, lookup, lookup_strong in *.
case_eq (M1 a). case_eq (M2 a).
intros; edestruct H; eauto.
auto.
destruct (M2 a); auto.
Qed.

Lemma disj_merge_disj :
  forall M1 M2 M3, disjoint M1 M3 -> disjoint M2 M3 -> disjoint (merge M1 M2) M3.
Proof.
 unfold disjoint. intros. eapply merge_indom in H1. firstorder.
Qed.

Lemma merge_assoc_L :
  forall M1 M2 M3, merge (merge M1 M2) M3 = merge M1 (merge M2 M3).
Proof.
  intros.
  apply functional_extensionality.
  intro.
  unfold merge.
  case_eq (M1 x); trivial.
Qed.

Lemma merge_assoc_R :
  forall M1 M2 M3, merge M1 (merge M2 M3) = merge (merge M1 M2) M3.
Proof.
intros. symmetry. apply merge_assoc_L.
Qed.

Lemma merge_disj_disj_1 :
  forall M1 M2 M3, disjoint (merge M1 M2) M3 -> disjoint M1 M3.
Proof.
unfold disjoint.
intros.
eauto using indom_merge_indom_1.
Qed.

Lemma merge_disj_disj_2 :
  forall M1 M2 M3, disjoint (merge M1 M2) M3 -> disjoint M2 M3.
Proof.
unfold disjoint.
intros.
eauto using indom_merge_indom_2.
Qed.


Definition dom (M : Map) (a : A): bool :=
  match M a with None => false | _ => true end.

Definition bapp (f g: A -> bool) (a: A): bool :=
  if f a then true else g a.

Lemma bapp_or : forall f g a, bapp f g a = true -> f a = true \/ g a = true.
Proof.
  unfold bapp.
  intros until a.
  case_eq (f a); auto.
Qed.

Lemma or_bapp: forall f g a, (f a = true \/ g a = true) -> bapp f g a = true.
Proof.
  unfold bapp.
  destruct 1.
   rewrite H. trivial.
  destruct (f a); auto.
Qed.

Lemma dom_merge: forall M1 M2, dom (merge M1 M2) = bapp (dom M1) (dom M2).
Proof.
  intros.
  apply functional_extensionality.
  intro.
  unfold dom, merge, bapp.
  destruct (M1 x); trivial.
Qed.

Lemma in_dom_dom : forall M a, in_dom M a -> dom M a = true.
Proof.
  unfold in_dom, dom, lookup, lookup_strong.
  destruct 1.
  rewrite H.
  trivial.
Qed.

Lemma dom_in_dom : forall M a, dom M a = true -> in_dom M a.
Proof.
  unfold in_dom, dom, lookup, lookup_strong.
  intros until a.
  destruct (M a).
   eauto.
  discriminate.
Qed.

Definition create_map (cdom : A -> bool) (f : A -> B) (a : A) : option B :=
  if cdom a then Some (f a) else None.

Lemma create_map_dom : forall cdom f, dom (create_map cdom f) = cdom.
Proof.
  unfold dom, create_map.
  intros.
  apply functional_extensionality.
  intros.
  destruct (cdom x); auto.
Qed.

Lemma create_map_lookup_intro : forall cdom a, cdom a = true -> forall f, lookup (create_map cdom f) a (f a).
Proof.
  unfold lookup, lookup_strong, create_map.
  intros.
  rewrite H.
  trivial.
Qed.

Lemma create_map_lookup_elim : forall cdom a f b, lookup (create_map cdom f) a b -> b = f a.
Proof.
  intros.
  refine (_ (create_map_lookup_intro _ f)).
  2: eapply in_dom_dom.
  2: esplit. 2: eassumption.
  rewrite create_map_dom.
  intros.
  eauto using lookup_func.
Qed.

Lemma create_map_ext: forall cdom f1 f2
  (Hf: forall a, cdom a = true -> f1 a = f2 a),
   (create_map cdom f1) = (create_map cdom f2).
Proof.
  intros.
  unfold create_map.
  apply functional_extensionality.
  intros.
  case_eq (cdom x).
  intros.
   f_equal.
   eauto.
  trivial.
Qed.

Lemma indom_eq:
forall (M M1:Map) a,
in_dom M a -> dom M = dom M1 -> in_dom M1 a.
Proof.
intros.
apply dom_in_dom.
rewrite <- H0.
apply in_dom_dom.
assumption.
Qed.

Lemma indom_merge_replace:
forall (M M1 M2:Map) a,
in_dom (merge M M1) a -> dom M1 = dom M2 -> in_dom (merge M M2) a.
Proof.
intros.
apply merge_indom in H.
destruct H.
apply indom_merge_indom_1.
assumption.
apply indom_merge_indom_2.
apply dom_in_dom.
rewrite <- H0.
apply in_dom_dom.
assumption.
Qed.

Lemma disjoint_same_dom: forall M1 M2, disjoint M1 M2 -> forall M3, dom M2 = dom M3 ->
  disjoint M1 M3.
Proof.
  unfold disjoint.
  intros.
  eapply H.
  eassumption.
  eapply dom_in_dom.
  rewrite H0.
  eapply in_dom_dom.
  assumption.
Qed.

Definition merge_embedded (l: list Map) : Map :=
  fold_right merge empty_map l.

Inductive disjoint_embedded: list Map -> Prop :=
| disjoint_embedded_nil:
  disjoint_embedded nil
| disjoint_embedded_cons: forall a l,
  disjoint a (merge_embedded l) ->
  disjoint_embedded l ->
  disjoint_embedded (a :: l)
.

End Map_values.

Variables B1 B2: Type.

Inductive same_domains_embedded : list (Map B1) -> list (Map B2) -> Prop :=
| same_domains_embedded_nil: same_domains_embedded nil nil
| same_domains_embedded_cons: forall m1 m2,
  dom m1 = dom m2 ->
  forall l1 l2,
  same_domains_embedded l1 l2 ->
  same_domains_embedded (m1 :: l1) (m2 :: l2).

Lemma disjoint_embedded_same_domains: forall l1,
  disjoint_embedded l1 ->
  forall l2, same_domains_embedded l1 l2 ->
    disjoint_embedded l2 /\ dom (merge_embedded l1) = dom (merge_embedded l2).
Proof.
  induction 1; inversion 1; subst.
   split.
   constructor.
   reflexivity.
  destruct (IHdisjoint_embedded _ H6).
  split.
   constructor.
   unfold disjoint.
   intros.
   eapply H.
   eapply dom_in_dom.
   rewrite H4.
   eapply in_dom_dom.
   eassumption.
   eapply dom_in_dom.
   rewrite H3.
   eapply in_dom_dom.
   eassumption.
   assumption.
  unfold merge_embedded in H3.
  unfold merge_embedded.
  simpl.
  rewrite dom_merge.
  rewrite dom_merge.
  rewrite H4.
  rewrite H3.
  trivial.
Qed.

End Map.

Opaque lookup_strong dom merge create_map.