From a8c744000247af207b489d3cdd4e3d3cf60f72e1 Mon Sep 17 00:00:00 2001 From: blazy Date: Fri, 8 Jan 2010 07:53:02 +0000 Subject: ajout branche allocation de registres git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@1220 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e --- backend/Affinity_relation.v | 152 + backend/Allocproof.v | 4 +- backend/Coloring.v | 19 +- backend/Coloringproof.v | 19 +- backend/Conservative_criteria.v | 339 + backend/Delete_Preference_Edges_Adjacency.v | 60 + backend/Delete_Preference_Edges_Degree.v | 16 + backend/Delete_Preference_Edges_Move.v | 138 + backend/Edges.v | 411 ++ backend/EqualSetMap.v | 82 + backend/Freeze_WL.v | 210 + backend/Graph_Facts.v | 191 + backend/Graph_translation.v | 3552 ++++++++++ backend/IRC.v | 122 + backend/IRCColoring.v | 847 +++ backend/IRC_Graph_Functions.v | 548 ++ backend/IRC_graph.v | 15 + backend/IRC_termination.v | 397 ++ backend/InterfGraph.v | 6 +- backend/InterfGraphMapImp.v | 9401 +++++++++++++++++++++++++++ backend/InterfGraphProperties.v | 8 + backend/InterfGraph_Construction.v | 189 + backend/Interference_adjacency.v | 66 + backend/Merge_Adjacency.v | 347 + backend/Merge_Degree.v | 205 + backend/Merge_Move.v | 730 +++ backend/Merge_WL.v | 818 +++ backend/MyAllocation.v | 1434 ++++ backend/MyRegisters.v | 143 + backend/Order_arith.v | 349 + backend/OrderedOption.v | 96 + backend/Regs.v | 122 + backend/Remove_Vertex_Adjacency.v | 73 + backend/Remove_Vertex_Degree.v | 173 + backend/Remove_Vertex_Move.v | 206 + backend/Remove_Vertex_WL.v | 478 ++ backend/SetsFacts.v | 105 + backend/Simplify_WL.v | 160 + backend/Spill_WL.v | 141 + backend/Typed_interfgraphs.v | 740 +++ backend/WS.v | 231 + 41 files changed, 23329 insertions(+), 14 deletions(-) create mode 100755 backend/Affinity_relation.v create mode 100755 backend/Conservative_criteria.v create mode 100755 backend/Delete_Preference_Edges_Adjacency.v create mode 100755 backend/Delete_Preference_Edges_Degree.v create mode 100755 backend/Delete_Preference_Edges_Move.v create mode 100755 backend/Edges.v create mode 100755 backend/EqualSetMap.v create mode 100755 backend/Freeze_WL.v create mode 100755 backend/Graph_Facts.v create mode 100644 backend/Graph_translation.v create mode 100755 backend/IRC.v create mode 100755 backend/IRCColoring.v create mode 100755 backend/IRC_Graph_Functions.v create mode 100755 backend/IRC_graph.v create mode 100755 backend/IRC_termination.v create mode 100755 backend/InterfGraphMapImp.v create mode 100755 backend/InterfGraphProperties.v create mode 100755 backend/InterfGraph_Construction.v create mode 100755 backend/Interference_adjacency.v create mode 100755 backend/Merge_Adjacency.v create mode 100755 backend/Merge_Degree.v create mode 100755 backend/Merge_Move.v create mode 100755 backend/Merge_WL.v create mode 100644 backend/MyAllocation.v create mode 100755 backend/MyRegisters.v create mode 100755 backend/Order_arith.v create mode 100755 backend/OrderedOption.v create mode 100755 backend/Regs.v create mode 100755 backend/Remove_Vertex_Adjacency.v create mode 100755 backend/Remove_Vertex_Degree.v create mode 100755 backend/Remove_Vertex_Move.v create mode 100755 backend/Remove_Vertex_WL.v create mode 100755 backend/SetsFacts.v create mode 100755 backend/Simplify_WL.v create mode 100755 backend/Spill_WL.v create mode 100755 backend/Typed_interfgraphs.v create mode 100755 backend/WS.v diff --git a/backend/Affinity_relation.v b/backend/Affinity_relation.v new file mode 100755 index 00000000..7c4029c9 --- /dev/null +++ b/backend/Affinity_relation.v @@ -0,0 +1,152 @@ +Require Import FSets. +Require Import InterfGraphMapImp. +Require Import Delete_Preference_Edges_Degree. +Require Import Edges. +Require Import MyRegisters. + +Module Register := Regs. + +Import Edge RegFacts Props. + +Lemma compat_bool_move : forall g, +compat_bool Register.eq (move_related g). + +Proof. +unfold move_related, compat_bool. intros. +rewrite (compat_preference_adj _ _ _ H). reflexivity. +Qed. + +(* Characterization of the move-relationship *) + +(* If a vertex x is move-related in g then there exists an affinity edge e + of g which is incident to x *) +Lemma move_related_charac : forall x g, +move_related g x = true -> +exists e, aff_edge e /\ In_graph_edge e g /\ incident e x. + +Proof. +intros. unfold move_related in H. +destruct (set_induction2 (preference_adj x g)). +rewrite Dec.F.is_empty_iff in H0. rewrite H0 in H. inversion H. +destruct H0. destruct H0. +rewrite Add_Equal in H0. +assert (VertexSet.In x0 (preference_adj x g)). +rewrite H0. apply VertexSet.add_1. intuition. +rewrite in_pref in H1. destruct H1. +exists (x0, x, Some x2). +split. unfold aff_edge. exists x2. auto. +split. assumption. +right. auto. +Qed. + +(* The inversion characterization of move related *) +Lemma move_related_charac2 : forall x g e, +aff_edge e -> +In_graph_edge e g -> +incident e x -> +move_related g x = true. + +Proof. +intros. unfold move_related. +case_eq (VertexSet.is_empty (preference_adj x g)); auto. +intros. rewrite <-Dec.F.is_empty_iff in H2. +generalize (empty_is_empty_1 H2). clear H2. intro. +destruct H. destruct H1. +assert (VertexSet.In (snd_ext e) (preference_adj x g)). +rewrite in_pref. exists x0. rewrite edge_comm. rewrite <-H. +assert (eq (x, snd_ext e, get_weight e) (fst_ext e, snd_ext e, get_weight e)) by Eq_eq. +rewrite H3. rewrite <-(edge_eq e). assumption. +rewrite H2 in H3. elim (VertexSet.empty_1 H3). +assert (VertexSet.In (fst_ext e) (preference_adj x g)). +rewrite in_pref. exists x0. rewrite <-H. +assert (eq (fst_ext e, x, get_weight e) (fst_ext e, snd_ext e, get_weight e)) by Eq_eq. +rewrite H3. rewrite <-(edge_eq e). assumption. +rewrite H2 in H3. elim (VertexSet.empty_1 H3). +Qed. + +(* Characterization of nonmove relation *) +Lemma move_related_false_charac : forall e x g, +aff_edge e -> +In_graph_edge e g -> +move_related g x = false -> +~incident e x. + +Proof. +intros e x g H H0 H1 H2. +generalize (move_related_charac2 _ _ _ H H0 H2). +intro H3;rewrite H1 in H3;inversion H3. +Qed. + +Lemma move_related_false_charac2 : forall x g, +(forall e, aff_edge e -> In_graph_edge e g -> ~incident e x) -> +move_related g x = false. + +Proof. +intros x g H. +case_eq (move_related g x);intro H0. +generalize (move_related_charac _ _ H0);intro H1. +destruct H1 as [e H1];destruct H1 as [H1 H2];destruct H2 as [H2 H3]. +elim ((H e H1 H2) H3). +reflexivity. +Qed. + +(* A move-related vertex of the graph belongs to the graph *) +Lemma move_related_in_graph : forall x g, +move_related g x = true -> In_graph x g. + +Proof. +intros x g H. +generalize (move_related_charac x g H);intro H0. +destruct H0 as [e H0];destruct H0 as [H0 H1];destruct H1 as [H1 H2]. +destruct H2;rewrite H2. +apply (proj1 (In_graph_edge_in_ext _ _ H1)). +apply (proj2 (In_graph_edge_in_ext _ _ H1)). +Qed. + +(* The endpoints of any affinity edge of g are move-related in g *) +Lemma Aff_edge_aff : forall e g, +In_graph_edge e g -> +aff_edge e -> +move_related g (fst_ext e) = true /\ move_related g (snd_ext e) = true. + +Proof. +intros. split. +apply move_related_charac2 with (e:=e); [idtac|idtac|left]; auto. +apply move_related_charac2 with (e:=e); [idtac|idtac|right]; auto. +Qed. + +(* Any move-related vertex has a nonempty preference neighborhood *) +Lemma move_related_not_empty_pref : forall x g, +move_related g x = true -> +~VertexSet.Equal (preference_adj x g) VertexSet.empty. + +Proof. +unfold move_related. intros. intro. +generalize (empty_is_empty_2 H0). intro. +rewrite Dec.F.is_empty_iff in H1. rewrite H1 in H. inversion H. +Qed. + +(* Any nonmove-related vertex has an empty preference neighborhood *) +Lemma not_move_related_empty_pref : forall x g, +move_related g x = false -> +VertexSet.Equal (preference_adj x g) VertexSet.empty. + +Proof. +unfold move_related. intros. +apply empty_is_empty_1. rewrite Dec.F.is_empty_iff. +case_eq (VertexSet.is_empty (preference_adj x g)); intros; auto. +rewrite H0 in H. inversion H. +Qed. + +(* Any vertex having a preference degree different than 0 is move-related *) +Lemma move_related_card : forall x g, +pref_degree g x <> 0 -> +move_related g x = true. + +Proof. +unfold pref_degree. intros. +case_eq (move_related g x); intros. +reflexivity. +generalize (not_move_related_empty_pref _ _ H0). intro. +rewrite H1 in H. rewrite empty_cardinal in H. elim H. auto. +Qed. diff --git a/backend/Allocproof.v b/backend/Allocproof.v index 10eaa5b1..c01b80c3 100644 --- a/backend/Allocproof.v +++ b/backend/Allocproof.v @@ -711,7 +711,9 @@ Proof. simpl. econstructor; eauto. change (transfer f (RTL.fn_entrypoint f) live !! (RTL.fn_entrypoint f)) with (live0 f live). - eapply agree_parameters; eauto. + eapply agree_parameters; eauto. + unfold regalloc. eauto. + unfold regalloc. congruence. (* external function *) injection H7; intro EQ; inv EQ. diff --git a/backend/Coloring.v b/backend/Coloring.v index 67824ae3..5282d480 100644 --- a/backend/Coloring.v +++ b/backend/Coloring.v @@ -22,6 +22,8 @@ Require Import RTLtyping. Require Import Locations. Require Import Conventions. Require Import InterfGraph. +Require Import MyAllocation. +Require Import InterfGraph_Construction. (** * Construction of the interference graph *) @@ -90,7 +92,7 @@ Require Import InterfGraph. - between the result of a ``call'' instruction and the location of the result as dictated by the calling conventions. *) - +(* Definition add_interf_live (filter: reg -> bool) (res: reg) (live: Regset.t) (g: graph): graph := Regset.fold @@ -209,10 +211,10 @@ Definition interf_graph (f: function) (live: PMap.t Regset.t) (live0: Regset.t) and the set of all [RTL] pseudo-registers mentioned in the interference graph. It returns the coloring as a function from pseudo-registers to locations. *) - +(* Parameter graph_coloring: function -> graph -> regenv -> Regset.t -> (reg -> loc). - +*) (** To ensure that the result of [graph_coloring] is a correct coloring, we check a posteriori its result using the following Coq functions. Let [coloring] be the function [reg -> loc] returned by [graph_coloring]. @@ -266,7 +268,7 @@ Definition check_coloring andb (check_coloring_1 g coloring) (andb (check_coloring_2 g coloring) (check_coloring_3 rs env coloring)). - +*) (** To preserve decidability of checking, the checks (especially the third one) are performed for the pseudo-registers mentioned in the interference graph. To facilitate the proofs, @@ -290,11 +292,14 @@ Definition alloc_of_coloring (coloring: reg -> loc) (env: regenv) (rs: Regset.t) and adjustment of this coloring. If the coloring candidate is incorrect, [None] is returned, causing register allocation to fail. *) +Definition graph_coloring (f : function) g env (rs : Regset.t) := +my_graph_coloring g env. + Definition regalloc (f: function) (live: PMap.t Regset.t) (live0: Regset.t) (env: regenv) := let g := interf_graph f live live0 in let rs := all_interf_regs g in let coloring := graph_coloring f g env rs in - if check_coloring g env rs coloring - then Some (alloc_of_coloring coloring env rs) - else None. +(* if check_coloring g env rs coloring *) +(* then *) Some (alloc_of_coloring coloring env rs). +(* else None. *) diff --git a/backend/Coloringproof.v b/backend/Coloringproof.v index 92ac0676..cc4234ec 100644 --- a/backend/Coloringproof.v +++ b/backend/Coloringproof.v @@ -24,6 +24,8 @@ Require Import Locations. Require Import Conventions. Require Import InterfGraph. Require Import Coloring. +Require Import InterfGraph_Construction. +Require Import MyAllocation. (** * Correctness of the interference graph *) @@ -446,12 +448,12 @@ Proof. intros until g2. intro. unfold correct_interf_instr; destruct instr; auto. destruct (is_move_operation o l). - intros. eapply interfere_incl; eauto. - intros. eapply interfere_incl; eauto. - intros. eapply interfere_incl; eauto. + intros. apply interfere_incl with (g1 := g1); auto. + intros. apply interfere_incl with (g1 := g1); auto. + intros. apply interfere_incl with (g1 := g1); auto. intros [A [B C]]. - split. intros. eapply interfere_mreg_incl; eauto. - split. intros. eapply interfere_incl; eauto. + split. intros. apply interfere_mreg_incl with (g1 := g1); auto. + split. intros. apply interfere_incl with (g1 := g1); auto. destruct s0; auto. intros. eapply interfere_mreg_incl; eauto. destruct s0; auto. intros. eapply interfere_mreg_incl; eauto. Qed. @@ -812,16 +814,23 @@ Let g := interf_graph f live live0. Let allregs := all_interf_regs g. Let coloring := graph_coloring f g env allregs. + Lemma regalloc_ok: regalloc f live live0 env = Some alloc -> check_coloring g env allregs coloring = true /\ alloc = alloc_of_coloring coloring env allregs. Proof. +unfold regalloc. intro. +inversion H. subst. clear H. +split. apply allocation_correct. +auto. +(* unfold regalloc, coloring, allregs, g. case (check_coloring (interf_graph f live live0) env). intro EQ; injection EQ; intro; clear EQ. split. auto. auto. intro; discriminate. +*) Qed. Lemma regalloc_acceptable: diff --git a/backend/Conservative_criteria.v b/backend/Conservative_criteria.v new file mode 100755 index 00000000..946a047b --- /dev/null +++ b/backend/Conservative_criteria.v @@ -0,0 +1,339 @@ +Require Import FSets. +Require Import InterfGraphMapImp. +Require Import Affinity_relation. +Require Import ZArith. +Require Import Edges. +Require Import MyRegisters. + +Import Edge RegFacts Props. + +(* Definition of the Briggs criterion *) + +Definition nb_of_significant_degree g K s := +VertexSet.fold (fun x n => if (le_lt_dec K (interf_degree g x)) then S n else n) + s 0. + +(* Definition of the conservative criteria : an affinity edge e is coalescible if + 1) one of its endpoints is not precolored + 2) the endpoints have less than K interference neighbors of high-degree *) +Definition conservative_coalescing_criteria e g K := +if (is_precolored (fst_ext e) g && is_precolored (snd_ext e) g) then false else +if le_lt_dec +K +(nb_of_significant_degree g K + (VertexSet.union (interference_adj (fst_ext e) g) + (interference_adj (snd_ext e) g))) +then false else true. + +(* Specification of the coalescing criterion *) +Lemma conservative_coalescing_criteria_1 : forall e g K, +~VertexSet.In (fst_ext e) (precolored g) \/ +~VertexSet.In (snd_ext e) (precolored g) -> +In_graph_edge e g -> +nb_of_significant_degree g K + (VertexSet.union (interference_adj (fst_ext e) g) (interference_adj (snd_ext e) g)) + < K -> +conservative_coalescing_criteria e g K = true. + +Proof. +intros e g K H HHHH H0. +assert (In_graph (fst_ext e) g /\ In_graph (snd_ext e) g). +apply In_graph_edge_in_ext. assumption. destruct H1 as [HH HHH]. +unfold conservative_coalescing_criteria. +case_eq (is_precolored (fst_ext e) g && is_precolored (snd_ext e) g); intros. +destruct H. +case_eq (is_precolored (fst_ext e) g); intros. +elim H. apply (precolored_equiv _ _). split; assumption. +rewrite H2 in H1. inversion H1. +case_eq (is_precolored (snd_ext e) g); intros. +destruct H. apply (precolored_equiv _ _). split; assumption. +rewrite H2 in H1. destruct (is_precolored (fst_ext e)); auto. +destruct (le_lt_dec K (nb_of_significant_degree g K + (VertexSet.union (interference_adj (fst_ext e) g) + (interference_adj (snd_ext e) g)))). +intuition. +reflexivity. +Qed. + +Lemma conservative_coalescing_criteria_2 : forall e g palette, +conservative_coalescing_criteria e g palette = true -> +~VertexSet.In (fst_ext e) (precolored g) \/ +~VertexSet.In (snd_ext e) (precolored g). + +Proof. +intros. +unfold conservative_coalescing_criteria in H. +case_eq (is_precolored (fst_ext e) g && is_precolored (snd_ext e) g); intros. +rewrite H0 in H. +inversion H. +rewrite H0 in H. +case_eq (is_precolored (fst_ext e) g);intros. +right. +case_eq (is_precolored (snd_ext e) g); intros. +rewrite H1 in H0. rewrite H2 in H0. inversion H0. +intro H3. generalize (proj1 (precolored_equiv _ _) H3). intro. +rewrite H2 in H4. destruct H4. inversion H4. +left. +intro H4. generalize (proj1 (precolored_equiv _ _) H4). intro. +rewrite H1 in H2. destruct H2. inversion H2. +Qed. + +Lemma conservative_coalescing_criteria_3 : forall e g K, +conservative_coalescing_criteria e g K = true -> +nb_of_significant_degree g K + (VertexSet.union (interference_adj (fst_ext e) g) (interference_adj (snd_ext e) g)) + < K. + +Proof. +intros. +unfold conservative_coalescing_criteria in H. +case_eq (is_precolored (fst_ext e) g && is_precolored (snd_ext e) g); intros. +rewrite H0 in H. inversion H. +rewrite H0 in H. +destruct (le_lt_dec + K + (nb_of_significant_degree g K + (VertexSet.union (interference_adj (fst_ext e) g) + (interference_adj (snd_ext e) g)))). +inversion H. +assumption. +Qed. + +Definition is_none (o : option Edge.t) := +match o with +|None => true +|Some _ => false +end. + +(* Function picking an edge satisfying a function f in a set s*) +Definition get_element_such_f (f : Edge.t -> bool) s := +EdgeSet.fold (fun e o => if (is_none o) then if (f e) then Some e else None else o) s None. + +(* Specification of the get_element_such function *) +Lemma element_some : forall l (f : Edge.t -> bool) a, +fold_left + (fun (a : option Edge.t) (e : EdgeSet.elt) => + if is_none a then if f e then Some e else None else a) l (Some a) = Some a. + +Proof. +induction l; simpl; intros. +reflexivity. +apply IHl. +Qed. + +Lemma get_element_correct : forall f s (x : Edge.t), +get_element_such_f f s = Some x -> f x = true. + +Proof. +unfold get_element_such_f. intros f s x H. +unfold get_element_such_f in H. +rewrite EdgeSet.fold_1 in H. +induction (EdgeSet.elements s); simpl in H. +simpl in H. inversion H. +case_eq (f a); intro H0; rewrite H0 in H. +rewrite element_some in H. inversion H. subst. assumption. +apply IHl. assumption. +Qed. + +Lemma get_element_in : forall f s x, +get_element_such_f f s = Some x -> EdgeSet.In x s. + +Proof. +intros f s x H. +unfold get_element_such_f in H. +rewrite EdgeSet.fold_1 in H. +generalize (EdgeSet.elements_2);intro H0. +generalize (H0 s);clear H0;intro H0. +induction (EdgeSet.elements s);simpl in *. +inversion H. +case_eq (f a); intro H1; rewrite H1 in H. +rewrite element_some in H. inversion H. subst. +apply H0. intuition. +apply IHl; intuition. +Qed. + +Lemma compat_is_precolored : forall x y g, +Register.eq x y -> +is_precolored x g = is_precolored y g. + +Proof. +exact is_precolored_ext. +Qed. + +(* Definition of the any_coalescible_edge function which picks an affinity + edge satisfying the coalescing criterion for K in g in the set moves *) +Definition any_coalescible_edge moves g K := +match (get_element_such_f (fun e => conservative_coalescing_criteria e g K) moves) with +|None => None +|Some e => if (is_precolored (fst_ext e) g) then Some e else Some (snd_ext e, fst_ext e, get_weight e) +end. + +Lemma adj_of_significant_degree_compat_set : forall g s s' K, +VertexSet.Equal s s' -> +nb_of_significant_degree g K s = +nb_of_significant_degree g K s'. + +Proof. +intros g s s' K H. +unfold nb_of_significant_degree. +cut (eqlistA Register.eq (VertexSet.elements s) (VertexSet.elements s')). intro H0. +do 2 rewrite VertexSet.fold_1. +generalize H0. generalize (VertexSet.elements s'). +clear H0. generalize 0. +induction (VertexSet.elements s). simpl. +intros n l H0. +destruct l. +simpl. reflexivity. +inversion H0. +simpl. intro n. +destruct l0. simpl. intro H0. inversion H0. +intro H0. simpl. +inversion H0. subst. +case_eq (le_lt_dec K (interf_degree g a)); intros H1 _; +case_eq (le_lt_dec K (interf_degree g e)); intros H2 _. +apply IHl. assumption. +unfold interf_degree in H1. rewrite (compat_interference_adj _ _ _ H4) in H1. +apply False_ind. intuition. +unfold interf_degree in H2. apply False_ind. intuition. +unfold interf_degree in H1. rewrite (compat_interference_adj _ _ _ H4) in H1. +unfold interf_degree in H2. apply False_ind. intuition. +apply IHl. assumption. +apply SortA_equivlistA_eqlistA with (ltA := Register.lt). +apply Register.eq_refl. +apply Register.eq_sym. +apply Register.eq_trans. +apply Register.lt_trans. +apply Register.lt_not_eq. +apply OTFacts.lt_eq. +apply OTFacts.eq_lt. +apply VertexSet.elements_3. +apply VertexSet.elements_3. +apply equal_equivlist. assumption. +Qed. + +Lemma compat_criteria_aux : forall e e' g K, +eq e e' -> +In_graph_edge e g -> +conservative_coalescing_criteria e g K = true -> +conservative_coalescing_criteria e' g K = true. + +Proof. +intros e e' g K H HH H0. +apply conservative_coalescing_criteria_1. +generalize (conservative_coalescing_criteria_2 _ _ _ H0). intro H1. +destruct H1. +destruct (eq_charac _ _ H); change_rewrite; destruct H2. +left. rewrite <-H2. assumption. +right. rewrite <-H2. assumption. +destruct (eq_charac _ _ H); change_rewrite; destruct H2. +right. rewrite <-H3. assumption. +left. rewrite <-H3. assumption. + +rewrite <-H. assumption. + +generalize (conservative_coalescing_criteria_3 _ _ _ H0). intro H1. +destruct (eq_charac _ _ H); change_rewrite; destruct H2 as [H2 H3]; +rewrite <-(compat_interference_adj _ _ _ H2) in *; +rewrite <-(compat_interference_adj _ _ _ H3) in *. +assumption. +rewrite adj_of_significant_degree_compat_set with +(s':= VertexSet.union (interference_adj (fst_ext e) g) (interference_adj (snd_ext e) g)). +assumption. +apply union_sym. +Qed. + +Lemma compat_criteria : forall e e' g K, +eq e e' -> +In_graph_edge e g -> +conservative_coalescing_criteria e g K = +conservative_coalescing_criteria e' g K. + +Proof. +intros e e' g palette HH H. +case_eq (conservative_coalescing_criteria e g palette); intros. +symmetry. apply compat_criteria_aux with (e:=e). assumption. assumption. assumption. +case_eq (conservative_coalescing_criteria e' g palette); intros. +assert (conservative_coalescing_criteria e g palette = true). +apply compat_criteria_aux with (e:=e'). apply eq_sym. assumption. +rewrite <-HH. assumption. assumption. +rewrite H0 in H2. inversion H2. +reflexivity. +Qed. + +(* Specification of any_coalescible_edge *) +Lemma any_coalescible_edge_1 : forall e g K s, +(forall e', EdgeSet.In e' s -> In_graph_edge e' g) -> +any_coalescible_edge s g K = Some e -> +conservative_coalescing_criteria e g K = true /\ EdgeSet.In e s. + +Proof. +intros e g K s HH H. +unfold any_coalescible_edge in H. +case_eq (get_element_such_f (fun e : t => conservative_coalescing_criteria e g K) s). +intros t0 H0. +rewrite H0 in H. +destruct (is_precolored (fst_ext t0) g);inversion H;subst. +split. +apply (get_element_correct _ _ _ H0). +apply (get_element_in _ _ _ H0). +assert (eq t0 (snd_ext t0, fst_ext t0, get_weight t0)) by (Eq_comm_eq; apply Regs.eq_refl). +split. +rewrite (compat_criteria _ _ _ _ (eq_sym H1)). +apply (get_element_correct _ _ _ H0). +apply HH. rewrite <-H1. apply (get_element_in _ _ _ H0). +rewrite <-H1. apply (get_element_in _ _ _ H0). +intro H0. rewrite H0 in H. inversion H. +Qed. + +Lemma any_coalescible_edge_11 : forall e g palette s, +any_coalescible_edge s g palette = Some e -> +EdgeSet.In e s. + +Proof. +intros. +unfold any_coalescible_edge in H. +case_eq (get_element_such_f + (fun e : t => conservative_coalescing_criteria e g palette) s); intros. +rewrite H0 in H. +case_eq (is_precolored (fst_ext t0) g); intros. +rewrite H1 in H. +inversion H. subst. +apply (get_element_in _ _ _ H0). +rewrite H1 in H. +inversion H. subst. +rewrite edge_comm. change Regs.registers with Regs.t. rewrite <-(edge_eq t0). + apply (get_element_in _ _ _ H0). +rewrite H0 in H. +inversion H. +Qed. + +Lemma any_coalescible_edge_2 : forall e g palette s, +(forall e', EdgeSet.In e' s -> In_graph_edge e' g) -> +any_coalescible_edge s g palette = Some e -> +~VertexSet.In (snd_ext e) (precolored g). + +Proof. +intros e g palette s H H0. +intro H1. +generalize (any_coalescible_edge_1 _ _ _ _ H H0). intro HH. +unfold any_coalescible_edge in H0. +case_eq (get_element_such_f (fun e : t =>conservative_coalescing_criteria e g palette) s). +intros t0 H2. +rewrite H2 in H0. +case_eq (is_precolored (fst_ext t0) g);intro H3. +rewrite H3 in H0. +generalize (proj1 (precolored_equiv _ _ ) H1);clear H1;intro H1. +inversion H0;subst. destruct HH as [HH H5]. +generalize (conservative_coalescing_criteria_2 _ _ _ HH). intro H4. +destruct H4; elim H4; apply (proj2 (precolored_equiv _ _)). +split. assumption. apply (proj1 (In_graph_edge_in_ext _ _ (H e (get_element_in _ _ _ H2)))). +destruct H1. split; assumption. +rewrite H3 in H0. +inversion H0;subst. +clear H0. +change (snd_ext (snd_ext t0,fst_ext t0,get_weight t0)) with (fst_ext t0) in H1. +generalize (proj1 (precolored_equiv _ _) H1);clear H1;intro H1. +rewrite H3 in H1;inversion H1. inversion H0. +intro H2;rewrite H2 in H0. +inversion H0. +Qed. diff --git a/backend/Delete_Preference_Edges_Adjacency.v b/backend/Delete_Preference_Edges_Adjacency.v new file mode 100755 index 00000000..64ace15a --- /dev/null +++ b/backend/Delete_Preference_Edges_Adjacency.v @@ -0,0 +1,60 @@ +Require Import FSets. +Require Import InterfGraphMapImp. +Require Import Merge_Degree. +Require Import Interference_adjacency. +Require Import Edges. +Require Import MyRegisters. + +Module Register := Regs. + +Import Edge RegFacts Props. + +(* The interference neighborhood of any vertex is left + unchanged when x is frozen *) +Lemma interf_adj_delete_preference : forall x r g H, +VertexSet.Equal (interference_adj x g) + (interference_adj x (delete_preference_edges r g H)). + +Proof. +split;intros. +rewrite in_interf. rewrite In_delete_preference_edges_edge. +rewrite in_interf in H0. split. assumption. +intro. destruct H1. +unfold aff_edge in H1. destruct H1. simpl in H1. congruence. + +rewrite in_interf in H0. rewrite In_delete_preference_edges_edge in H0. destruct H0. +rewrite in_interf. assumption. +Qed. + +(* The preference neighborhood of any vertex different from r + is obtained by removing r from its preference neighborhood in g *) +Lemma delete_preference_preference_adj : forall x r g H, +~Register.eq x r -> +VertexSet.Equal (preference_adj x (delete_preference_edges r g H)) + (VertexSet.remove r (preference_adj x g)). + +Proof. +intros. +split; intros. +apply VertexSet.remove_2. +intro. rewrite <-H2 in H1. +generalize (pref_adj_comm _ _ _ H1). intro. +rewrite in_pref in H3. destruct H3. +rewrite In_delete_preference_edges_edge in H3. +destruct H3. elim H4. split. +unfold aff_edge. exists x0. auto. +right. auto. +rewrite in_pref in H1. +destruct H1. rewrite In_delete_preference_edges_edge in H1. +destruct H1. rewrite in_pref. exists x0. assumption. + +(* <= *) +assert (~Register.eq r a). +intro. elim (VertexSet.remove_1 H2 H1). +generalize (VertexSet.remove_3 H1). clear H1. intro. +rewrite in_pref in H1. destruct H1. +rewrite in_pref. exists x0. +rewrite In_delete_preference_edges_edge. split. assumption. +intro. destruct H3. +destruct H4; change_rewrite. elim (H2 H4). elim H0; auto. +Qed. diff --git a/backend/Delete_Preference_Edges_Degree.v b/backend/Delete_Preference_Edges_Degree.v new file mode 100755 index 00000000..6b781f22 --- /dev/null +++ b/backend/Delete_Preference_Edges_Degree.v @@ -0,0 +1,16 @@ +Require Import FSets. +Require Import InterfGraphMapImp. +Require Import Delete_Preference_Edges_Adjacency. +Require Import Edges. + +Import Edge Props RegFacts. + +(* The interference degree is left unchanged when r is frozen. Hence, + a vertex is of low-degree after freezing r iff it is before freezing r *) +Lemma delete_preference_edges_low : forall x r g K p, +has_low_degree g K x = has_low_degree (delete_preference_edges r g p) K x. + +Proof. +intros x r g K p. unfold has_low_degree, interf_degree. +rewrite <-(Equal_cardinal (interf_adj_delete_preference x r g p)). reflexivity. +Qed. diff --git a/backend/Delete_Preference_Edges_Move.v b/backend/Delete_Preference_Edges_Move.v new file mode 100755 index 00000000..66f9ab6a --- /dev/null +++ b/backend/Delete_Preference_Edges_Move.v @@ -0,0 +1,138 @@ +Require Import FSets. +Require Import InterfGraphMapImp. +Require Import Merge_Move. +Require Import Edges. +Require Import MyRegisters. +Require Import Affinity_relation. +Require Import Interference_adjacency. +Require Import Delete_Preference_Edges_Adjacency. +Require Import Remove_Vertex_Move. + +Import Edge Props RegFacts. + +(* The frozen vertex is nonmove-related in the resulting graph *) +Lemma not_aff_related_delete_preference_edges : forall r g p, +move_related (delete_preference_edges r g p) r = false. + +Proof. +intros. apply move_related_false_charac2. +intros. +rewrite In_delete_preference_edges_edge in H0. destruct H0. +intro. elim H1. split; assumption. +Qed. + +(* A vertex which is move-related after a freeze is move-related before it *) +Lemma move_related_delete_move_related : forall x r g p, +move_related (delete_preference_edges r g p) x = true -> +move_related g x = true. + +Proof. +intros x r g p H. +generalize (move_related_charac _ _ H);intro H0. +destruct H0 as [y H0]. +destruct H0 as [H0 H1];destruct H1 as [H1 H2]. +rewrite In_delete_preference_edges_edge in H1. destruct H1. +apply (move_related_charac2 _ _ _ H0 H1 H2). +Qed. + +(* A vertex which is move-related in g and not in (delete_preference_edges r g H) + is a preference neighbor of r in g *) +Lemma delete_preference_edges_move_1 : forall x r g H, +~Register.eq r x -> +move_related g x = true -> +move_related (delete_preference_edges r g H) x = false -> +VertexSet.In x (preference_adj r g). + +Proof. +intros. generalize (not_move_related_empty_pref _ _ H2). intro. +rewrite delete_preference_preference_adj in H3. +destruct (In_dec r (preference_adj x g)). +apply pref_adj_comm. assumption. +assert (VertexSet.Equal (preference_adj x g) VertexSet.empty). +split; intros. +rewrite <-H3. apply VertexSet.remove_2. +intro. rewrite <-H5 in H4. elim n. auto. +assumption. +elim (VertexSet.empty_1 H4). +elim (move_related_not_empty_pref _ _ H1 H4). +auto. +Qed. + +(* A vertex which is move-related in g and nonmove-related in + (delete_preference_edges r g H) has a preference degree of 1 *) +Lemma delete_preference_edges_move_2 : forall x r g H, +~Register.eq r x -> +move_related g x = true -> +move_related (delete_preference_edges r g H) x = false -> +VertexSet.cardinal (preference_adj x g) = 1. + +Proof. +intros. generalize (not_move_related_empty_pref _ _ H2). intro. +rewrite delete_preference_preference_adj in H3. +destruct (In_dec r (preference_adj x g)). +cut (VertexSet.Equal (preference_adj x g) (VertexSet.singleton r)). intros. +rewrite H4. apply singleton_cardinal. +split; intros. +destruct (Register.eq_dec a r). +apply VertexSet.singleton_2. intuition. +assert (VertexSet.In a VertexSet.empty). +rewrite <-H3. apply VertexSet.remove_2; auto. +elim (VertexSet.empty_1 H5). +generalize (VertexSet.singleton_1 H4). intro. +rewrite <-H5. auto. +elim (move_related_not_empty_pref _ _ H1). +rewrite <-H3. + +split; intros. +rewrite Dec.F.remove_neq_iff. +assumption. +intro. rewrite <-H5 in H4. elim n. auto. +apply (VertexSet.remove_3 H4). +auto. +Qed. + +(* Reciprocally, a vertex which is move-related in g, + has a preference degree of 1 and is a preference neighbor of r + is nonmove-related in (delete_preference_edges r g H) *) +Lemma delete_preference_edges_move_inv : forall x r g H, +~Register.eq r x -> +move_related g x = true -> +VertexSet.In x (preference_adj r g) -> +VertexSet.cardinal (preference_adj x g) = 1 -> +move_related (delete_preference_edges r g H) x = false. + +Proof. +intros. +apply move_related_false_charac2. intros. +rewrite In_delete_preference_edges_edge in H5. +destruct H5. +intro. elim H6. +split. assumption. +cut (VertexSet.Equal (preference_adj x g) (VertexSet.singleton r)). intros. +destruct H7;[right|left]. +rewrite (edge_eq e) in H5. +assert (eq (x, snd_ext e, get_weight e) (fst_ext e, snd_ext e, get_weight e)). +Eq_eq. rewrite <-H9 in H5. +destruct H4. rewrite H4 in H5. +apply VertexSet.singleton_1. rewrite <-H8. +rewrite in_pref. exists x0. rewrite edge_comm. auto. +rewrite (edge_eq e) in H5. +assert (eq (fst_ext e, x, get_weight e) (fst_ext e, snd_ext e, get_weight e)). +Eq_eq. rewrite <-H9 in H5. +destruct H4. rewrite H4 in H5. +apply VertexSet.singleton_1. rewrite <-H8. +rewrite in_pref. exists x0. rewrite edge_comm. auto. +rewrite edge_comm. assumption. +apply cardinal_1_singleton. apply pref_adj_comm. assumption. assumption. +Qed. + +Lemma delete_preference_edges_move_false_false : forall x r g H, +move_related g x = false -> +move_related (delete_preference_edges r g H) x = false. + +Proof. +intros. +case_eq (move_related (delete_preference_edges r g H) x); intros. +rewrite (move_related_delete_move_related _ _ _ _ H1) in H0. inversion H0. +reflexivity. +Qed. diff --git a/backend/Edges.v b/backend/Edges.v new file mode 100755 index 00000000..1b570454 --- /dev/null +++ b/backend/Edges.v @@ -0,0 +1,411 @@ +Require Import FSets. +Require Import OrderedOption. +Require Import ZArith. +Require Import MyRegisters. + +(* Module of edges in a simple graph : there is never more + than one edge between two vertices *) + +Module Edge. + +Module Import Vertex := Regs. + +(* Definition of pair of vertices, to represent edges *) +Module VertexPair := PairOrderedType Vertex Vertex. + +(* N_as_OT is the OrderedType with t := N, + to describe weights of edges *) +Module OptionN_as_OT := OrderedOpt N_as_OT. + +(* An edge is finally a pair of endpoints and a weight *) +Module E := PairOrderedType VertexPair OptionN_as_OT. + +(* Useful modules imports *) +Module Import OTFacts := OrderedTypeFacts Vertex. +Module OTPairFacts := OrderedTypeFacts VertexPair. +Module OTEFacts := OrderedTypeFacts E. + +(* t is simply E.t *) +Definition t := E.t. + +(* accessors to the edges, their endpoints and their weight *) +Definition get_edge : t -> (Vertex.t*Vertex.t) := fun x => fst x. +Definition fst_ext : t -> Vertex.t := fun x => fst (get_edge x). +Definition snd_ext : t -> Vertex.t := fun x => snd (get_edge x). +Definition get_weight : t -> option N := fun x => snd x. + +(* get_edge e is only the extremities of e *) +Lemma get_edge_ext : forall e, +get_edge e = (fst_ext e,snd_ext e). + +Proof. +unfold fst_ext. unfold snd_ext. unfold get_edge. +simpl. intro e. destruct (fst e);auto. +Qed. + +(* An edge is the pair of vertices and a weight *) +Lemma edge_edge_weight : forall e, +e = (get_edge e, get_weight e). + +Proof. +unfold get_edge. unfold get_weight. +simpl. intro e. destruct e;auto. +Qed. + +(* Expansion of an edge *) +Lemma edge_eq : forall e, +e = (fst_ext e, snd_ext e, get_weight e). + +Proof. +intro e. unfold get_weight. +destruct e. unfold fst_ext. unfold snd_ext. unfold get_edge. simpl. +destruct p. auto. +Qed. + +(* Equality does not depend on the order of endpoints, e.g + (x,y,w) is equal to (y,x,w) so we need to define ordered edges *) +Definition ordered_edge e := +match (lt_dec (snd_ext e) (fst_ext e)) with +|left _ => (snd_ext e, fst_ext e, get_weight e) +|right _ => e +end. + +(* Commutativity of ordered_edge*) +Lemma ordered_edge_comm : forall x y o, +E.eq (ordered_edge (x,y,o)) (ordered_edge (y,x,o)). + +Proof. +unfold ordered_edge, snd_ext, fst_ext, get_edge. intros x y o. simpl. +destruct (lt_dec y x);destruct (lt_dec x y). +elim (lt_not_gt r r0). +apply E.eq_refl. +apply E.eq_refl. +destruct (Vertex.compare x y). +elim (n0 l). +unfold E.eq. simpl. repeat split; auto. +apply Vertex.eq_sym. auto. +apply OptionN_as_OT.eq_refl. +elim (n l). +Qed. + +(* Definition of equality as equality of s *) +Definition eq x y := E.eq (ordered_edge x) (ordered_edge y). + +(* The weak equality is the equality of their type (interference or preference) + and of their extremities, but not necessarily of their weight *) +Definition weak_eq x y := +(Vertex.eq (fst_ext x) (fst_ext y) /\ Vertex.eq (snd_ext x) (snd_ext y)) \/ +(Vertex.eq (fst_ext x) (snd_ext y) /\ Vertex.eq (snd_ext x) (fst_ext y)). + +(* Two edges having equal extremities (in the right order) and equal weights + are equal *) +Lemma eq_ordered_eq : forall x y, +E.eq x y -> eq x y. + +Proof. +unfold eq;unfold E.eq;intros x y H. +unfold ordered_edge;unfold fst_ext;unfold snd_ext;unfold get_edge;simpl. +destruct x as [x wx];destruct x as [x1 x2]; +destruct y as [y wy];destruct y as [y1 y2];simpl in *. +destruct (lt_dec x2 x1); destruct (lt_dec y2 y1);simpl in *. +intuition. +destruct H as [H HH];destruct H as [H H0]. +elim n. apply eq_lt with (y := x2). + apply Vertex.eq_sym;assumption. + apply lt_eq with (y := x1);assumption. +destruct H as [H HH];destruct H as [H H0]. +elim n. apply eq_lt with (y := y2). + assumption. + apply lt_eq with (y := y1); auto. apply Regs.eq_sym. auto. +assumption. +Qed. + +(* Definition of lt as the lexicographic order on endpoints + of the edges, after ordering *) +Definition lt x y := VertexPair.lt (get_edge (ordered_edge x)) (get_edge (ordered_edge y)) \/ + (VertexPair.eq (get_edge (ordered_edge x)) (get_edge (ordered_edge y)) /\ + OptionN_as_OT.lt (get_weight x) (get_weight y)). + +Lemma eq_refl : forall x, eq x x. + +Proof. +unfold eq;unfold E.eq. +repeat split;[apply Vertex.eq_refl|apply Vertex.eq_refl|apply OptionN_as_OT.eq_refl]. +Qed. + +Lemma eq_sym : forall x y, eq x y -> eq y x. + +Proof. +unfold eq;unfold E.eq;intros x y H. intuition. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +apply OptionN_as_OT.eq_sym;assumption. +Qed. + +Lemma eq_trans : forall x y z, eq x y -> eq y z -> eq x z. + +Proof. +unfold eq;intros x y z H H0. +apply (E.eq_trans _ _ _ H H0). +Qed. + +Lemma lt_trans : forall x y z, lt x y -> lt y z -> lt x z. + +Proof. +unfold lt;unfold get_edge;unfold ordered_edge;intros x y z H H0. +destruct (lt_dec (snd_ext x) (fst_ext x));simpl in *. +destruct (lt_dec (snd_ext y) (fst_ext y));simpl in *. +destruct (lt_dec (snd_ext z) (fst_ext z));simpl in *. +destruct H;destruct H0. +left;apply (VertexPair.lt_trans _ _ _ H H0). +left;apply OTPairFacts.lt_eq with (y := (snd_ext y, fst_ext y));simpl. +unfold VertexPair.lt in H;simpl in H;assumption. +destruct H0 as [H0 HH0]. +unfold VertexPair.eq in H0;assumption. +left;apply OTPairFacts.eq_lt with (y := (snd_ext y, fst_ext y));simpl. +destruct H as [H HH]. +unfold VertexPair.eq in H;unfold VertexPair.lt in H0; +simpl in H0;assumption. +unfold VertexPair.lt in H0;simpl in H0;assumption. +right;destruct H as [H HH];destruct H0 as [H0 HH0];split. +apply (VertexPair.eq_trans _ _ _ H H0). +apply (OptionN_as_OT.lt_trans _ _ _ HH HH0). +destruct H;destruct H0. +left;apply (VertexPair.lt_trans _ _ _ H H0). +left;apply OTPairFacts.lt_eq with (y := (snd_ext y, fst_ext y));simpl. +unfold VertexPair.lt in H;simpl in H;assumption. +destruct H0 as [H0 HH0]. +unfold VertexPair.eq in H0;assumption. +left;apply OTPairFacts.eq_lt with (y := (snd_ext y, fst_ext y));simpl. +destruct H as [H HH]. +unfold VertexPair.eq in H;unfold VertexPair.lt in H0; +simpl in H0;assumption. +unfold VertexPair.lt in H0;simpl in H0;assumption. +right;destruct H as [H HH];destruct H0 as [H0 HH0];split. +apply (VertexPair.eq_trans _ _ _ H H0). +apply (OptionN_as_OT.lt_trans _ _ _ HH HH0). +destruct (lt_dec (snd_ext y) (fst_ext y));simpl in *. +elim (n r0). +destruct (lt_dec (snd_ext z) (fst_ext z));simpl in *. +destruct H;destruct H0. +left;apply (VertexPair.lt_trans _ _ _ H H0). +left;apply OTPairFacts.lt_eq with (y := (fst y));simpl. +unfold VertexPair.lt in H;simpl in H;assumption. +destruct H0 as [H0 HH0]. +unfold VertexPair.eq in H0;assumption. +left;apply OTPairFacts.eq_lt with (y := (fst y));simpl. +destruct H as [H HH]. +unfold VertexPair.eq in H;unfold VertexPair.lt in H0; +simpl in H0;assumption. +unfold VertexPair.lt in H0;simpl in H0;assumption. +right;destruct H as [H HH];destruct H0 as [H0 HH0];split. +apply (VertexPair.eq_trans _ _ _ H H0). +apply (OptionN_as_OT.lt_trans _ _ _ HH HH0). +destruct H;destruct H0. +left;apply (VertexPair.lt_trans _ _ _ H H0). +left;apply OTPairFacts.lt_eq with (y := (fst y));simpl. +unfold VertexPair.lt in H;simpl in H;assumption. +destruct H0 as [H0 HH0]. +unfold VertexPair.eq in H0;assumption. +left;apply OTPairFacts.eq_lt with (y := (fst y));simpl. +destruct H as [H HH]. +unfold VertexPair.eq in H;unfold VertexPair.lt in H0; +simpl in H0;assumption. +unfold VertexPair.lt in H0;simpl in H0;assumption. +right;destruct H as [H HH];destruct H0 as [H0 HH0];split. +apply (VertexPair.eq_trans _ _ _ H H0). +apply (OptionN_as_OT.lt_trans _ _ _ HH HH0). +destruct (lt_dec (snd_ext y) (fst_ext y));simpl in *. +destruct (lt_dec (snd_ext z) (fst_ext z));simpl in *. +destruct H;destruct H0. +left;apply (VertexPair.lt_trans _ _ _ H H0). +left;apply OTPairFacts.lt_eq with (y := (snd_ext y, fst_ext y));simpl. +unfold VertexPair.lt in H;simpl in H;assumption. +destruct H0 as [H0 HH0]. +unfold VertexPair.eq in H0;assumption. +left;apply OTPairFacts.eq_lt with (y := (snd_ext y,fst_ext y));simpl. +destruct H as [H HH]. +unfold VertexPair.eq in H;unfold VertexPair.lt in H0; +simpl in H0;assumption. +unfold VertexPair.lt in H0;simpl in H0;assumption. +right;destruct H as [H HH];destruct H0 as [H0 HH0];split. +apply (VertexPair.eq_trans _ _ _ H H0). +apply (OptionN_as_OT.lt_trans _ _ _ HH HH0). +destruct H;destruct H0. +left;apply (VertexPair.lt_trans _ _ _ H H0). +left;apply OTPairFacts.lt_eq with (y := (snd_ext y,fst_ext y));simpl. +unfold VertexPair.lt in H;simpl in H;assumption. +destruct H0 as [H0 HH0]. +unfold VertexPair.eq in H0;assumption. +left;apply OTPairFacts.eq_lt with (y := (snd_ext y, fst_ext y));simpl. +destruct H as [H HH]. +unfold VertexPair.eq in H;unfold VertexPair.lt in H0; +simpl in H0;assumption. +unfold VertexPair.lt in H0;simpl in H0;assumption. +right;destruct H as [H HH];destruct H0 as [H0 HH0];split. +apply (VertexPair.eq_trans _ _ _ H H0). +apply (OptionN_as_OT.lt_trans _ _ _ HH HH0). +destruct (lt_dec (snd_ext y) (fst_ext y));simpl in *. +elim (n0 r). +destruct (lt_dec (snd_ext z) (fst_ext z));simpl in *. +destruct H;destruct H0. +left;apply (VertexPair.lt_trans _ _ _ H H0). +left;apply OTPairFacts.lt_eq with (y := (fst y));simpl. +unfold VertexPair.lt in H;simpl in H;assumption. +destruct H0 as [H0 HH0]. +unfold VertexPair.eq in H0;assumption. +left;apply OTPairFacts.eq_lt with (y := (fst y));simpl. +destruct H as [H HH]. +unfold VertexPair.eq in H;unfold VertexPair.lt in H0; +simpl in H0;assumption. +unfold VertexPair.lt in H0;simpl in H0;assumption. +right;destruct H as [H HH];destruct H0 as [H0 HH0];split. +apply (VertexPair.eq_trans _ _ _ H H0). +apply (OptionN_as_OT.lt_trans _ _ _ HH HH0). +destruct H;destruct H0. +left;apply (VertexPair.lt_trans _ _ _ H H0). +left;apply OTPairFacts.lt_eq with (y := (fst y));simpl. +unfold VertexPair.lt in H;simpl in H;assumption. +destruct H0 as [H0 HH0]. +unfold VertexPair.eq in H0;assumption. +left;apply OTPairFacts.eq_lt with (y := (fst y));simpl. +destruct H as [H HH]. +unfold VertexPair.eq in H;unfold VertexPair.lt in H0; +simpl in H0;assumption. +unfold VertexPair.lt in H0;simpl in H0;assumption. +right;destruct H as [H HH];destruct H0 as [H0 HH0];split. +apply (VertexPair.eq_trans _ _ _ H H0). +apply (OptionN_as_OT.lt_trans _ _ _ HH HH0). +Qed. + +Lemma weight_ordered_weight : forall x, +get_weight x = get_weight (ordered_edge x). + +Proof. +intro x;unfold ordered_edge;unfold get_weight;simpl. +destruct (lt_dec (snd_ext x) (fst_ext x));simpl;reflexivity. +Qed. + +Lemma lt_not_eq : forall x y, lt x y -> ~eq x y. + +Proof. +unfold lt;unfold eq;unfold E.eq;intros x y H. +destruct H. +generalize (VertexPair.lt_not_eq _ _ H);intro H0. +unfold VertexPair.eq in H0. +intro H1; elim H0. +destruct H1 as [H1 H2]. +unfold get_edge;assumption. +destruct H as [H H0]. +unfold VertexPair.eq in H. +unfold get_edge in H. +intro H1. +destruct H1 as [H1 HH1]. +elim (OptionN_as_OT.lt_not_eq _ _ H0). +rewrite (weight_ordered_weight x). +rewrite (weight_ordered_weight y). +unfold get_weight;assumption. +Qed. + +Lemma compare : forall x y, Compare lt eq x y. + +Proof. +intros x y. +destruct (OTEFacts.lt_dec (ordered_edge x) (ordered_edge y)). +apply LT. +unfold lt;unfold get_edge;unfold VertexPair.lt; +unfold VertexPair.eq. +rewrite (weight_ordered_weight x). +rewrite (weight_ordered_weight y). +unfold get_weight. +destruct (ordered_edge x) as [ex wx];destruct ex as [ex1 ex2]; +destruct (ordered_edge y) as [ey wy];destruct ey as [ey1 ey2];simpl in *. +assumption. +destruct (OTEFacts.eq_dec (ordered_edge x) (ordered_edge y)). +apply EQ. +unfold eq. +destruct (ordered_edge x) as [ex wx];destruct ex as [ex1 ex2]; +destruct (ordered_edge y) as [ey wy];destruct ey as [ey1 ey2];simpl in *. +destruct a as [H H0];destruct H as [H H1]. +unfold E.eq;auto. +apply GT. +generalize (OTEFacts.le_neq n n0);intro H. +unfold lt;unfold get_edge;unfold VertexPair.lt; +unfold VertexPair.eq. +rewrite (weight_ordered_weight x). +rewrite (weight_ordered_weight y). +unfold get_weight. +destruct (ordered_edge x) as [ex wx];destruct ex as [ex1 ex2]; +destruct (ordered_edge y) as [ey wy];destruct ey as [ey1 ey2];simpl in *. +assumption. +Qed. + +(* Edges commutativity *) +Lemma edge_comm : forall x y o, eq (x,y,o) (y,x,o). + +Proof. +unfold eq;apply ordered_edge_comm. +Qed. + +(* Definition of affinity edges *) +Definition aff_edge x := exists w, get_weight x = Some w. + +(* Definition of interference edges *) +Definition interf_edge x := get_weight x = None. + +(* An edge is incident to a vertex iff this vertex is an endpoint of the edge *) +Definition incident e x := Vertex.eq x (fst_ext e) \/ Vertex.eq x (snd_ext e). + +(* An edge is incident to a vertex, or is not *) +Lemma incident_dec : forall e x, +incident e x \/ ~incident e x. + +Proof. +intros e x. +destruct (eq_dec x (fst_ext e)). +left;unfold incident;left;assumption. +destruct (eq_dec x (snd_ext e)). +left;unfold incident;right;assumption. +right;unfold incident;intro Heq. +destruct Heq;[elim (n H)|elim (n0 H)]. +Qed. + +(* Definition redirect : redirect x y e replaces the extremity x of the edge e + with y, if e is incident to x *) +Definition redirect x y e := +match eq_dec (fst_ext e) x with +|left _ => (y, snd_ext e, get_weight e) +|right _ => match eq_dec (snd_ext e) x with + | left _ => (fst_ext e, y, get_weight e) + | right _ => e + end +end. + +(* Decidable equality over edges *) +Lemma eq_dec : forall x y, {eq x y}+{~eq x y}. + +Proof. +intros x y. +destruct (compare x y). +right. apply lt_not_eq. assumption. +left. assumption. +right. intro H. generalize (eq_sym _ _ H). intro H0. elim (lt_not_eq _ _ l H0). +Qed. + +(* Equality of edges implies constraints of their endpoints *) +Lemma eq_charac : forall x y, +eq x y -> (Vertex.eq (fst_ext x) (fst_ext y) /\ Vertex.eq (snd_ext x) (snd_ext y)) \/ + (Vertex.eq (fst_ext x) (snd_ext y) /\ Vertex.eq (snd_ext x) (fst_ext y)). + +Proof. +intros x y H;unfold eq in H;unfold ordered_edge in H; +unfold get_edge in H. +destruct (lt_dec (snd_ext x) (fst_ext x)); +destruct (lt_dec (snd_ext y) (fst_ext y)); +unfold E.eq in H;simpl in H;intuition. +Qed. + +Definition same_type e e' := (aff_edge e /\ aff_edge e') \/ + (interf_edge e /\ interf_edge e'). + +End Edge. \ No newline at end of file diff --git a/backend/EqualSetMap.v b/backend/EqualSetMap.v new file mode 100755 index 00000000..292b1c0e --- /dev/null +++ b/backend/EqualSetMap.v @@ -0,0 +1,82 @@ +Require Import List. +Require Import InterfGraphMapImp. + +Section fold_assoc_interf_map. + +Variable A : Type. + +Inductive eq_set_option : option VertexSet.t -> option VertexSet.t -> Prop := +|None_eq : eq_set_option None None +|Some_eq : forall s s', VertexSet.Equal s s' -> eq_set_option (Some s) (Some s'). + +Definition EqualSetMap m1 m2 := forall x, eq_set_option (VertexMap.find x m1) (VertexMap.find x m2). + +Lemma EqualSetMap_refl : forall m, EqualSetMap m m. + +Proof. +unfold EqualSetMap. intro m. intro x. +destruct (VertexMap.find x m). +constructor. intuition. +constructor. +Qed. + +Lemma EqualSetMap_trans : forall m1 m2 m3, +EqualSetMap m1 m2 -> +EqualSetMap m2 m3 -> +EqualSetMap m1 m3. + +Proof. +intros m1 m2 m3 H H0. +unfold EqualSetMap in *. +intro x. +generalize (H x). clear H. intro H. +generalize (H0 x). clear H0. intro H0. +destruct (VertexMap.find x m1). +inversion H. subst. +rewrite <-H2 in H0. +destruct (VertexMap.find x m3). +constructor. inversion H0. subst. +rewrite H3. assumption. +inversion H0. +destruct (VertexMap.find x m3). +inversion H0. inversion H. subst. rewrite <-H4 in H1. inversion H1. +constructor. +Qed. + +Lemma fold_left_compat_map : forall (f : VertexMap.t VertexSet.t -> A -> VertexMap.t VertexSet.t) l e e', +EqualSetMap e e' -> +(forall e1 e2 a, EqualSetMap e1 e2 -> EqualSetMap (f e1 a) (f e2 a)) -> +EqualSetMap (fold_left f l e) (fold_left f l e'). + +Proof. +intro f;induction l;simpl. +auto. +intros e e' H H0 H1. +apply (IHl (f e a) (f e' a)). +apply H0;assumption. +assumption. +Qed. + +Lemma fold_left_assoc_map : forall l (f : VertexMap.t VertexSet.t -> A -> VertexMap.t VertexSet.t) x h, +(forall (y z : A) s, EqualSetMap (f (f s y) z) (f (f s z) y)) -> +(forall e1 e2 a, EqualSetMap e1 e2 -> EqualSetMap (f e1 a) (f e2 a)) -> +EqualSetMap (fold_left f (h :: l) x) (f (fold_left f l x) h). + +Proof. +induction l;simpl;intros f x h H H0. +apply EqualSetMap_refl. +apply EqualSetMap_trans with (m2 := fold_left f (h :: l) (f x a)). +simpl. apply fold_left_compat_map. apply H. +assumption. +apply IHl. assumption. assumption. +Qed. +(* +Add Morphism EqualSetMap : equalsetmap_m. + +Proof. +unfold EqualSetMap, VertexMap.Equal. split; intros. +rewrite <-(H0 x1). rewrite <-(H x1). apply H1. +rewrite (H0 x1). rewrite (H x1). apply H1. +Qed. +*) +End fold_assoc_interf_map. \ No newline at end of file diff --git a/backend/Freeze_WL.v b/backend/Freeze_WL.v new file mode 100755 index 00000000..a235374b --- /dev/null +++ b/backend/Freeze_WL.v @@ -0,0 +1,210 @@ +Require Import FSets. +Require Import InterfGraphMapImp. +Require Import Merge_WL. +Require Import Edges. +Require Import Remove_Vertex_WL. +Require Import IRC_graph. +Require Import Delete_Preference_Edges_Degree. +Require Import WS. +Require Import Delete_Preference_Edges_Move. +Require Import Graph_Facts. +Require Import Interference_adjacency. +Require Import Affinity_relation. + +Import RegFacts Props. + +Definition delete_preferences_wl2 v ircg k := +let wl := irc_wl ircg in +let g := irc_g ircg in +let simplify := get_simplifyWL wl in +let freeze := get_freezeWL wl in +let spill := get_spillWL wl in +let moves := get_movesWL wl in +let adj_fst := VertexSet.diff (preference_adj v g) (precolored g) in +let new_simp := VertexSet.filter (fun x => eq_K 1 (VertexSet.cardinal (preference_adj x g)) && has_low_degree g k x) adj_fst in +let simplifyWL' := VertexSet.add v (VertexSet.union simplify new_simp) in +let freezeWL' := VertexSet.remove v (VertexSet.diff freeze new_simp) in +let movesWL' := not_incident_edges v moves g in +(spill, freezeWL', simplifyWL', movesWL'). + +Lemma WS_delete_preferences_wl2 : forall r ircg (Hin : In_graph r (irc_g ircg)), +VertexSet.In r (get_freezeWL (irc_wl ircg)) -> +WS_properties (delete_preference_edges r (irc_g ircg) Hin) + (VertexSet.cardinal (pal ircg)) + (delete_preferences_wl2 r ircg (VertexSet.cardinal (pal ircg))). + + +Proof. +intros r ircg Hin HH. +unfold WS_properties. unfold delete_preferences_wl2. +generalize (HWS_irc ircg). intro HWS. +set (wl := irc_wl ircg) in *. +set (g := irc_g ircg) in *. +set (simplify := get_simplifyWL wl) in *. +set (freeze := get_freezeWL wl) in *. +set (spill := get_spillWL wl) in *. +set (adj_fst := VertexSet.diff (preference_adj r g) (precolored g)) in *. +set (new_simp := VertexSet.filter (fun x => eq_K 1 (VertexSet.cardinal (preference_adj x g))) adj_fst) in *. +set (simplifyWL_ := VertexSet.union simplify new_simp) in *. +set (simplifyWL' := VertexSet.add r simplifyWL_) in *. +set (freezeWL' := VertexSet.diff freeze new_simp) in *. +set (movesWL' := not_incident_edges r (get_movesWL wl) g) in *. +rewrite <-(Hk ircg) in HWS. set (K := VertexSet.cardinal (pal ircg)) in *. + +unfold get_spillWL, get_simplifyWL, get_freezeWL, get_movesWL. simpl. + +split. intro x. +rewrite <-delete_preference_edges_low. +rewrite precolored_delete_preference_edges. +rewrite In_delete_preference_edges_vertex. +split; intros. +apply (In_spill_props _ _ _ _ _ _ _ _ H (refl_equal _) HWS). +WS_apply HWS. assumption. + +(* freeze *) +split. intro x. +rewrite precolored_delete_preference_edges. +rewrite <-delete_preference_edges_low. + +split; intros. +unfold freezeWL' in H. +assert (~Register.eq r x) as Hneq. +intro. elim (VertexSet.remove_1 H0 H). +generalize (VertexSet.remove_3 H). clear H. intro H. +generalize (VertexSet.diff_1 H). generalize (VertexSet.diff_2 H). clear H. intros. +generalize (In_freeze_props _ _ _ _ _ _ _ _ H0 (refl_equal _) HWS). intro. +destruct H1. destruct H2. destruct H3. +split. +assumption. +split. + +destruct (Register.eq_dec r x). elim (Hneq e). +case_eq (move_related (delete_preference_edges r g Hin) x); intros. +reflexivity. +elim H. unfold new_simp. +apply VertexSet.filter_3. +apply compat_move_up. +apply VertexSet.diff_3. +apply delete_preference_edges_move_1 with (H := Hin); auto. +assumption. +rewrite andb_true_iff. split. +apply eq_K_1. +apply delete_preference_edges_move_2 with (H := Hin); auto. +assumption. +assumption. + +destruct H. destruct H0. +apply VertexSet.remove_2. intro. +rewrite (compat_bool_move _ _ _ (Register.eq_sym _ _ H2)) in H0. +rewrite (not_aff_related_delete_preference_edges r g Hin) in H0. inversion H0. +unfold freezeWL'. apply VertexSet.diff_3. +WS_apply HWS. split. +assumption. +split. +apply (move_related_delete_move_related _ _ _ _ H0). +assumption. +intro. assert (move_related (delete_preference_edges r g Hin) x = false). +apply delete_preference_edges_move_inv. +intro. +rewrite (compat_bool_move _ _ _ (Register.eq_sym _ _ H3)) in H0. +rewrite (not_aff_related_delete_preference_edges r g Hin) in H0. inversion H0. +apply (move_related_delete_move_related _ _ _ _ H0). +unfold new_simp in H2. +apply (VertexSet.diff_1 (VertexSet.filter_1 (compat_move_up _ _) H2)). +symmetry. apply eq_K_2. generalize (VertexSet.filter_2 (compat_move_up _ _) H2). intro. +rewrite andb_true_iff in H3. destruct H3. assumption. +rewrite H0 in H3. inversion H3. + +(* simplify *) +split. intro x. +rewrite <-delete_preference_edges_low. +rewrite precolored_delete_preference_edges. +rewrite In_delete_preference_edges_vertex. +split; intros. +unfold simplifyWL' in H. +destruct (proj1 (Dec.F.add_iff _ _ _) H). +generalize (In_freeze_props _ _ _ _ _ _ _ _ HH (refl_equal _) HWS). intro. +destruct H1. destruct H2. destruct H3. +split. rewrite <-(compat_bool_low _ _ _ _ H0). assumption. +split. rewrite <-(compat_bool_move _ _ _ H0). apply not_aff_related_delete_preference_edges. +split. rewrite <-H0. assumption. +rewrite <-H0. assumption. + +unfold simplifyWL_ in H0. +destruct (VertexSet.union_1 H0). +generalize (In_simplify_props _ _ _ _ _ _ _ _ H1 (refl_equal _) HWS). intro. +destruct H2. destruct H3. destruct H4. +split. +assumption. +split. +apply delete_preference_edges_move_false_false. assumption. +split; assumption. +unfold new_simp in H1. +generalize (VertexSet.filter_1 (compat_move_up _ _) H1). intro. +generalize (VertexSet.filter_2 (compat_move_up _ _) H1). clear H1. intro. +rewrite andb_true_iff in H1. destruct H1. +split. assumption. +split. apply delete_preference_edges_move_inv. +intro. elim (not_in_pref_self r g). +rewrite <-H4 in H2. apply (VertexSet.diff_1 H2). +apply move_related_card. +unfold pref_degree. rewrite <-(eq_K_2 _ _ H1). auto. +apply (VertexSet.diff_1 H2). +symmetry. apply (eq_K_2 _ _ H1). +split. +apply (in_pref_in _ _ _ (VertexSet.diff_1 H2)). +apply (VertexSet.diff_2 H2). + +destruct H. destruct H0. destruct H1. +destruct (Register.eq_dec r x). +apply VertexSet.add_1. auto. +apply VertexSet.add_2. +case_eq (In_dec x adj_fst); intros. +case_eq (eq_K 1 (VertexSet.cardinal (preference_adj x g))); intros. +apply VertexSet.union_3. apply VertexSet.filter_3. +apply compat_move_up. +assumption. +rewrite H4. rewrite H. auto. +apply VertexSet.union_2. WS_apply HWS. +split. assumption. +split. +case_eq (move_related g x); intros. +generalize (delete_preference_edges_move_2 _ _ _ Hin n H5 H0). intro. +rewrite (eq_K_1 _ _ H6) in H4. inversion H4. +reflexivity. +split; assumption. +apply VertexSet.union_2. +WS_apply HWS. +split. assumption. +split. case_eq (move_related g x); intros. +generalize (delete_preference_edges_move_1 _ _ _ Hin n H4 H0). intro. +elim n0. apply VertexSet.diff_3; assumption. +reflexivity. +split; assumption. + +unfold movesWL'. intro. rewrite not_incident_edges_1. +rewrite In_delete_preference_edges_edge. +split; intros. +destruct H. +generalize (In_move_props _ _ _ _ _ _ _ _ H (refl_equal _) HWS). intro. +destruct H1. +split. assumption. +split. assumption. +intro. elim H0. destruct H3. assumption. +destruct H. destruct H0. +split. WS_apply HWS. split; assumption. +intro. elim H1. split; assumption. + +intros. apply (In_move_props _ _ _ _ _ _ _ _ H (refl_equal _) HWS). +Qed. + +Lemma WS_freeze : forall r ircg (Hin : In_graph r (irc_g ircg)), +VertexSet.In r (get_freezeWL (irc_wl ircg)) -> +WS_properties (delete_preference_edges r (irc_g ircg) Hin) + (irc_k ircg) + (delete_preferences_wl2 r ircg (irc_k ircg)). + + +Proof. +intros. rewrite <-(Hk ircg). apply WS_delete_preferences_wl2. auto. +Qed. diff --git a/backend/Graph_Facts.v b/backend/Graph_Facts.v new file mode 100755 index 00000000..3919b1f3 --- /dev/null +++ b/backend/Graph_Facts.v @@ -0,0 +1,191 @@ +Require Import FSets. +Require Import InterfGraphMapImp. +Require Import ZArith. +Require Import Edges. +Require Import MyRegisters. + +Import Edge OTFacts. +Module Import RegFacts := Facts VertexSet. +Module Import RegRegFacts := Facts EdgeSet. + +(* Definition of useful morphisms *) + +Add Morphism In_graph : In_graph_m. + +Proof. +unfold In_graph;intros x y H g. +split;intro H0;[rewrite <-H|rewrite H];assumption. +Qed. + +Add Morphism In_graph_edge : In_graph_edge_m. + +Proof. +unfold In_graph_edge; intros x y H g. fold (eq x y) in H. +split;intro H0;[rewrite <-H|rewrite H];assumption. +Qed. + +Add Morphism move_related : move_related_m. + +Proof. +unfold move_related. intros. rewrite (compat_preference_adj _ _ _ H). reflexivity. +Qed. + +Add Morphism get_weight : get_weight_m. + +Proof. +intros x y H. +rewrite (weight_ordered_weight x);rewrite (weight_ordered_weight y). +unfold get_weight;unfold E.eq in H. +destruct H as [_ H];inversion H;[|rewrite H2];reflexivity. +Qed. + +(* Useful rewriting lemmas *) + +Lemma rewrite_fst : forall x y z, +fst_ext (x,y,z) = x. + +Proof. +auto. +Qed. + +Lemma rewrite_snd : forall x y z, +snd_ext (x,y,z) = y. + +Proof. +auto. +Qed. + +Lemma rewrite_weight : forall x y z, +get_weight (x,y,z) = z. + +Proof. +auto. +Qed. + +(* A rewriting tactic *) + +Ltac change_rewrite := +repeat (try rewrite rewrite_fst in *; + try rewrite rewrite_snd in *; + try rewrite rewrite_weight in *). + +(* Two tactics for proving equality of edges *) +Ltac Eq_eq := +apply eq_ordered_eq;unfold E.eq; +split;[simpl;split;auto; intuition |try apply OptionN_as_OT.eq_refl; intuition]. + +Ltac Eq_comm_eq := rewrite edge_comm; Eq_eq. + +(* Extensionnality of redirect *) +Lemma redirect_ext : forall e e' x y, +eq e e' -> +eq (redirect x y e) (redirect x y e'). + +Proof. +intros e e' x y H. +destruct e as [xe we];destruct xe as [ex1 ex2]; +destruct e' as [xe' we'];destruct xe' as [e'x1 e'x2];simpl in *. +generalize (get_weight_m _ _ H);simpl;intro Heq;subst. +destruct (eq_charac _ _ H);destruct H0 as [H0 H1];unfold redirect;change_rewrite. +destruct (OTFacts.eq_dec ex1 x). +destruct (OTFacts.eq_dec e'x1 x). +Eq_eq. intuition. +elim (n (Regs.eq_trans (Regs.eq_sym H0) r)). +destruct (OTFacts.eq_dec e'x1 x). +elim (n (Regs.eq_trans H0 r)). +destruct (OTFacts.eq_dec ex2 x). +destruct (OTFacts.eq_dec e'x2 x). +Eq_eq. intuition. +elim (n1 (Regs.eq_trans (Regs.eq_sym H1) r)). +destruct (OTFacts.eq_dec e'x2 x). +elim (n1 (Regs.eq_trans H1 r)). +Eq_eq. +destruct (OTFacts.eq_dec ex1 x). +destruct (OTFacts.eq_dec e'x1 x). +Eq_eq. intuition. +apply (Regs.eq_trans H1 (Regs.eq_trans r0 (Regs.eq_trans (Regs.eq_sym r) H0))). +destruct (OTFacts.eq_dec e'x2 x). +Eq_comm_eq. intuition. +elim (n0 (Regs.eq_trans (Regs.eq_sym H0) r)). +destruct (OTFacts.eq_dec e'x2 x). +elim (n (Regs.eq_trans H0 r)). +destruct (OTFacts.eq_dec ex2 x). +destruct (OTFacts.eq_dec e'x1 x). +Eq_comm_eq. +elim (n1 (Regs.eq_trans (Regs.eq_sym H1) r)). +destruct (OTFacts.eq_dec e'x1 x). +elim (n1 (Regs.eq_trans H1 r)). +Eq_comm_eq. +Qed. + +(* The weight is left unchanged while applying redirect *) +Lemma redirect_weight_eq : forall e x y, +Edge.get_weight (redirect x y e) = Edge.get_weight e. + +Proof. +unfold redirect;intros e x y;destruct e as [e w];destruct e as [ex1 ex2];change_rewrite. +destruct (OTFacts.eq_dec ex1 x);destruct (OTFacts.eq_dec ex2 x);reflexivity. +Qed. + +(* Specification of redirect *) +Lemma redirect_charac : forall e x y, +(eq e (redirect x y e) /\ (~Regs.eq x (fst_ext e) /\ ~Regs.eq x (snd_ext e))) \/ +(eq (y, snd_ext e, get_weight e) (redirect x y e) /\ Regs.eq x (fst_ext e)) \/ +(eq (fst_ext e, y, get_weight e) (redirect x y e) /\ Regs.eq x (snd_ext e)). + +Proof. +intros e x y. +unfold redirect. change_rewrite. +destruct (OTFacts.eq_dec (fst_ext e) x). +right. left. split. apply eq_refl. auto. +destruct (OTFacts.eq_dec (snd_ext e) x). +right. right. split. apply eq_refl. auto. +left. split. apply eq_refl. split; intuition. +Qed. + +(* Weak equality implies equality for interference edges *) +Lemma weak_eq_interf_eq : forall x y, +weak_eq x y -> +interf_edge x -> +interf_edge y -> +eq x y. + +Proof. +unfold weak_eq, interf_edge, get_weight. intros. destruct H; destruct H. +Eq_eq. rewrite H0. rewrite H1. apply OptionN_as_OT.eq_refl. +rewrite (edge_eq x). rewrite (edge_eq y). +Eq_comm_eq. simpl. unfold get_weight. +rewrite H0. rewrite H1. apply OptionN_as_OT.eq_refl. +Qed. + +(* The second endpoint of e does not belongs to (merge e g H H0) *) +Lemma not_in_merge_snd : forall e g H H0, +~In_graph (snd_ext e) (merge e g H H0). + +Proof. +intros. intro. +rewrite In_merge_vertex in H1. destruct H1. +elim (H2 (Regs.eq_refl _)). +Qed. + +(* Extensionnality of the low-degree function *) +Lemma compat_bool_low : forall g K, +compat_bool Regs.eq (has_low_degree g K). + +Proof. +unfold compat_bool, has_low_degree, interf_degree. intros. +rewrite (compat_interference_adj _ _ _ H). reflexivity. +Qed. + +(* There cannot exist both an interference and + a preference between two vertices *) +Lemma interf_pref_conflict : forall x y g, +Prefere x y g /\ Interfere x y g -> False. + +Proof. +unfold Prefere, Interfere. intros. destruct H. destruct H. +assert (eq (x,y,Some x0) (x,y,None)). +apply is_simple_graph with (g:=g); auto. +unfold weak_eq. left. change_rewrite. split; auto. +generalize (get_weight_m _ _ H1). simpl. congruence. +Qed. diff --git a/backend/Graph_translation.v b/backend/Graph_translation.v new file mode 100644 index 00000000..fb8f6c22 --- /dev/null +++ b/backend/Graph_translation.v @@ -0,0 +1,3552 @@ +Require Import InterfGraph. +Require Import InterfGraphMapImp. +Require Import MyRegisters. +Require Import RTLtyping. +Require Import NArith. +Require Import List. +Require Import InterfGraph_Construction. +Require Import Edges. +Require Import EqualSetMap. + +Import Edge. + +Definition new_graph := tt. + +Section Translation. + +Variable g : graph. +Variable interfgraph_vertices : VertexSet.t. + +Definition IE_reg_reg := +SetRegReg.fold (fun e s => EdgeSet.add (Regs.reg_to_Reg (fst e), Regs.reg_to_Reg (snd e), None) s) + (interf_reg_reg g) EdgeSet.empty. + +Definition IE_reg_mreg := +SetRegMreg.fold (fun e s => EdgeSet.add (Regs.reg_to_Reg (fst e), Regs.mreg_to_Reg (snd e), None) s) + (interf_reg_mreg g) EdgeSet.empty. + +Definition AE_reg_reg := +SetRegReg.fold (fun e s => EdgeSet.add (Regs.reg_to_Reg (fst e), Regs.reg_to_Reg (snd e), Some N0) s) + (pref_reg_reg g) EdgeSet.empty. + +Definition AE_reg_mreg := +SetRegMreg.fold (fun e s => EdgeSet.add (Regs.reg_to_Reg (fst e), Regs.mreg_to_Reg (snd e), Some N0) s) + (pref_reg_mreg g) EdgeSet.empty. + +Definition interfgraph_interference_edges := EdgeSet.union IE_reg_reg IE_reg_mreg. + +Definition resolve_conflicts s1 s2 := +EdgeSet.fold (fun e s => if EdgeSet.mem (fst_ext e, snd_ext e,None) s2 + then s else EdgeSet.add e s) s1 EdgeSet.empty. + +Definition interfgraph_affinity_edges := +resolve_conflicts +(EdgeSet.union AE_reg_reg AE_reg_mreg) +interfgraph_interference_edges. + +Hypothesis extremities_interf_graph : +forall e, EdgeSet.In e interfgraph_affinity_edges \/ + EdgeSet.In e interfgraph_interference_edges -> + VertexSet.In (fst_ext e) interfgraph_vertices /\ + VertexSet.In (snd_ext e) interfgraph_vertices. + +Definition new_adj_set x y m := +match (VertexMap.find x m) with + | None => VertexSet.singleton y + | Some s => VertexSet.add y s +end. + +Definition empty_map := +VertexSet.fold + (fun x m => VertexMap.add x VertexSet.empty m) + interfgraph_vertices + (VertexMap.empty VertexSet.t). + +Definition set2map s map := +EdgeSet.fold (fun e m => VertexMap.add (fst_ext e) (new_adj_set (fst_ext e) (snd_ext e) m) + (VertexMap.add (snd_ext e) (new_adj_set (snd_ext e) (fst_ext e) m) m)) + s map. + +Definition interf_map := set2map interfgraph_interference_edges empty_map. +Definition pref_map := set2map interfgraph_affinity_edges empty_map. + +Lemma add_simpl : forall l a s s2, +EdgeSet.Equal + (fold_left + (fun (a0 : EdgeSet.t) (e0 : EdgeSet.elt) => + if EdgeSet.mem (fst_ext e0, snd_ext e0, None) s2 + then a0 + else EdgeSet.add e0 a0) l (EdgeSet.add a s)) +(EdgeSet.add a +(fold_left + (fun (a0 : EdgeSet.t) (e0 : EdgeSet.elt) => + if EdgeSet.mem (fst_ext e0, snd_ext e0, None) s2 + then a0 + else EdgeSet.add e0 a0) l s)). + +Proof. +induction l;simpl;intros. +intuition. +case_eq (EdgeSet.mem (fst_ext a,snd_ext a,None) s2);intro H. +apply IHl. +do 3 rewrite IHl. +apply RegRegProps.add_add. +Qed. + + +Lemma resolve_conflicts_2 : forall e s1 s2, +EdgeSet.In e (resolve_conflicts s1 s2) -> +EdgeSet.In e s1 /\ ~EdgeSet.In (fst_ext e,snd_ext e, None) s2. + +Proof. +intros e s1 s2 H. +unfold resolve_conflicts in H. +rewrite EdgeSet.fold_1 in H. +split. +generalize (EdgeSet.elements_2);intro H0;generalize (H0 s1);clear H0;intro H0. +induction (EdgeSet.elements s1);simpl in *. +elim (EdgeSet.empty_1 H). +destruct (eq_dec a e). +apply H0. +fold (eq a e) in e0. +generalize (Edge.eq_sym e0);clear e0;intro e0. +intuition. +case_eq (EdgeSet.mem (fst_ext a, snd_ext a,None) s2);intro H1;rewrite H1 in H. +apply IHl. +apply H. +intuition. +apply IHl. + +rewrite add_simpl in H. +apply (EdgeSet.add_3 n H). +intuition. +induction (EdgeSet.elements s1);simpl in *. +elim (EdgeSet.empty_1 H). +case_eq (EdgeSet.mem (fst_ext a,snd_ext a,None) s2);intro H0;rewrite H0 in H. +apply IHl;assumption. +destruct (eq_dec a e). +intro H1. +destruct (eq_charac _ _ e0). +destruct H2 as [H2 HH2]. +assert (eq (fst_ext a,snd_ext a,None) (fst_ext e,snd_ext e,None)) by Eq_eq. +rewrite (RegRegProps.Dec.F.mem_b s2 H3) in H0. +generalize (EdgeSet.mem_1 H1);clear H1;intro H1. +rewrite H0 in H1;inversion H1. +destruct H2 as [H2 HH2]. +assert (eq (fst_ext a,snd_ext a,None) (fst_ext e,snd_ext e,None)) by Eq_comm_eq. +rewrite (RegRegProps.Dec.F.mem_b s2 H3) in H0. +generalize (EdgeSet.mem_1 H1);clear H1;intro H1. +rewrite H0 in H1;inversion H1. +apply IHl. +rewrite add_simpl in H. +apply (EdgeSet.add_3 n H). +Qed. + +(* +intros g e H. +unfold interfgraph_vertices. +destruct H. +unfold interfgraph_affinity_edges in H. +generalize (proj1 (resolve_conflicts_2 _ _ _ H)). clear H. intro H. +destruct (EdgeSet.union_1 H). +cut (VertexSet.In (fst_ext e) (V_pref_reg_reg g) /\ + VertexSet.In (snd_ext e) (V_pref_reg_reg g)). +intro H1. destruct H1 as [H1 H2]. split; +(do 2 apply VertexSet.union_3; +apply VertexSet.union_2; assumption). +unfold V_pref_reg_reg. +unfold AE_reg_reg in H0. +rewrite SetRegReg.fold_1 in *. +induction (SetRegReg.elements (pref_reg_reg g)). +simpl in H0. elim (EdgeSet.empty_1 H0). +rewrite MEdgeFacts.fold_left_assoc in H0. +destruct (proj1 (RegRegProps.Dec.F.add_iff _ _ _) H0). +fold (eq (Regs.reg_to_Reg (fst a), Regs.reg_to_Reg (snd a), Some N0) e) in H1. +rewrite RegFacts.fold_left_assoc. +destruct (eq_charac _ _ H1); destruct H2 as [H2 H3]. +change_rewrite. split. +apply VertexSet.add_1. assumption. +apply VertexSet.add_2. apply VertexSet.add_1. assumption. +change_rewrite. split. +apply VertexSet.add_2. apply VertexSet.add_1. assumption. +apply VertexSet.add_1. assumption. +intros y z s. +rewrite (Props.add_add (VertexSet.add (Regs.reg_to_Reg (snd y)) s) + (Regs.reg_to_Reg (snd z)) + (Regs.reg_to_Reg (fst y))). +rewrite (Props.add_add s + (Regs.reg_to_Reg (snd z)) + (Regs.reg_to_Reg (snd y))). +set (s' := VertexSet.add (Regs.reg_to_Reg (snd z) ) s). +rewrite (Props.add_add s' + (Regs.reg_to_Reg (snd y)) + (Regs.reg_to_Reg (fst z))). +apply Props.add_add. +intros. apply Props.Dec.F.add_m. +apply Regs.eq_refl. +apply Props.Dec.F.add_m;[apply Regs.eq_refl|assumption]. +rewrite RegFacts.fold_left_assoc. +split; do 2 apply VertexSet.add_2. +apply (proj1 (IHl H1)). +apply (proj2 (IHl H1)). +intros y z s. +rewrite (Props.add_add (VertexSet.add (Regs.reg_to_Reg (snd y)) s) + (Regs.reg_to_Reg (snd z)) + (Regs.reg_to_Reg (fst y))). +rewrite (Props.add_add s + (Regs.reg_to_Reg (snd z)) + (Regs.reg_to_Reg (snd y))). +set (s' := VertexSet.add (Regs.reg_to_Reg (snd z) ) s). +rewrite (Props.add_add s' + (Regs.reg_to_Reg (snd y)) + (Regs.reg_to_Reg (fst z))). +apply Props.add_add. +intros. apply Props.Dec.F.add_m. +apply Regs.eq_refl. +apply Props.Dec.F.add_m;[apply Regs.eq_refl|assumption]. +intros;apply RegRegProps.add_add. +intros;apply RegRegProps.Dec.F.add_m;[apply eq_refl|assumption]. + +cut (VertexSet.In (fst_ext e) (V_pref_reg_mreg g) /\ + VertexSet.In (snd_ext e) (V_pref_reg_mreg g)). +intro H1. destruct H1 as [H1 H2]. split; +(do 3 apply VertexSet.union_3; assumption). +unfold V_pref_reg_mreg. +unfold AE_reg_mreg in H0. +rewrite SetRegMreg.fold_1 in *. +induction (SetRegMreg.elements (pref_reg_mreg g)). +simpl in H0. elim (EdgeSet.empty_1 H0). +rewrite MEdgeFacts.fold_left_assoc in H0. +destruct (proj1 (RegRegProps.Dec.F.add_iff _ _ _) H0). +fold (eq (Regs.reg_to_Reg (fst a), Regs.mreg_to_Reg (snd a), Some N0) e) in H1. +rewrite RegFacts.fold_left_assoc. +destruct (eq_charac _ _ H1); destruct H2 as [H2 H3]. +change_rewrite. split. +apply VertexSet.add_1. assumption. +apply VertexSet.add_2. apply VertexSet.add_1. assumption. +change_rewrite. split. +apply VertexSet.add_2. apply VertexSet.add_1. assumption. +apply VertexSet.add_1. assumption. +intros y z s. +rewrite (Props.add_add (VertexSet.add (Regs.mreg_to_Reg (snd y)) s) + (Regs.mreg_to_Reg (snd z)) + (Regs.reg_to_Reg (fst y))). +rewrite (Props.add_add s + (Regs.mreg_to_Reg (snd z)) + (Regs.mreg_to_Reg (snd y))). +set (s' := VertexSet.add (Regs.mreg_to_Reg (snd z) ) s). +rewrite (Props.add_add s' + (Regs.mreg_to_Reg (snd y)) + (Regs.reg_to_Reg (fst z))). +apply Props.add_add. +intros. apply Props.Dec.F.add_m. +apply Regs.eq_refl. +apply Props.Dec.F.add_m;[apply Regs.eq_refl|assumption]. +rewrite RegFacts.fold_left_assoc. +split; do 2 apply VertexSet.add_2. +apply (proj1 (IHl H1)). +apply (proj2 (IHl H1)). +intros y z s. +rewrite (Props.add_add (VertexSet.add (Regs.mreg_to_Reg (snd y)) s) + (Regs.mreg_to_Reg (snd z)) + (Regs.reg_to_Reg (fst y))). +rewrite (Props.add_add s + (Regs.mreg_to_Reg (snd z)) + (Regs.mreg_to_Reg (snd y))). +set (s' := VertexSet.add (Regs.mreg_to_Reg (snd z) ) s). +rewrite (Props.add_add s' + (Regs.mreg_to_Reg (snd y)) + (Regs.reg_to_Reg (fst z))). +apply Props.add_add. +intros. apply Props.Dec.F.add_m. +apply Regs.eq_refl. +apply Props.Dec.F.add_m;[apply Regs.eq_refl|assumption]. +intros;apply RegRegProps.add_add. +intros;apply RegRegProps.Dec.F.add_m;[apply eq_refl|assumption]. + +unfold interfgraph_interference_edges in H. +destruct (EdgeSet.union_1 H). +cut (VertexSet.In (fst_ext e) (V_interf_reg_reg g) /\ + VertexSet.In (snd_ext e) (V_interf_reg_reg g)). +intro H1. destruct H1 as [H1 H2]. split; +(apply VertexSet.union_2; assumption). +unfold V_interf_reg_reg. +unfold IE_reg_reg in H0. +rewrite SetRegReg.fold_1 in *. +induction (SetRegReg.elements (interf_reg_reg g)). +simpl in H0. elim (EdgeSet.empty_1 H0). +rewrite MEdgeFacts.fold_left_assoc in H0. +destruct (proj1 (RegRegProps.Dec.F.add_iff _ _ _) H0). +fold (eq (Regs.reg_to_Reg (fst a), Regs.reg_to_Reg (snd a), Some N0) e) in H1. +rewrite RegFacts.fold_left_assoc. +destruct (eq_charac _ _ H1); destruct H2 as [H2 H3]. +change_rewrite. split. +apply VertexSet.add_1. assumption. +apply VertexSet.add_2. apply VertexSet.add_1. assumption. +change_rewrite. split. +apply VertexSet.add_2. apply VertexSet.add_1. assumption. +apply VertexSet.add_1. assumption. +intros y z s. +rewrite (Props.add_add (VertexSet.add (Regs.reg_to_Reg (snd y)) s) + (Regs.reg_to_Reg (snd z)) + (Regs.reg_to_Reg (fst y))). +rewrite (Props.add_add s + (Regs.reg_to_Reg (snd z)) + (Regs.reg_to_Reg (snd y))). +set (s' := VertexSet.add (Regs.reg_to_Reg (snd z) ) s). +rewrite (Props.add_add s' + (Regs.reg_to_Reg (snd y)) + (Regs.reg_to_Reg (fst z))). +apply Props.add_add. +intros. apply Props.Dec.F.add_m. +apply Regs.eq_refl. +apply Props.Dec.F.add_m;[apply Regs.eq_refl|assumption]. +rewrite RegFacts.fold_left_assoc. +split; do 2 apply VertexSet.add_2. +apply (proj1 (IHl H1)). +apply (proj2 (IHl H1)). +intros y z s. +rewrite (Props.add_add (VertexSet.add (Regs.reg_to_Reg (snd y)) s) + (Regs.reg_to_Reg (snd z)) + (Regs.reg_to_Reg (fst y))). +rewrite (Props.add_add s + (Regs.reg_to_Reg (snd z)) + (Regs.reg_to_Reg (snd y))). +set (s' := VertexSet.add (Regs.reg_to_Reg (snd z) ) s). +rewrite (Props.add_add s' + (Regs.reg_to_Reg (snd y)) + (Regs.reg_to_Reg (fst z))). +apply Props.add_add. +intros. apply Props.Dec.F.add_m. +apply Regs.eq_refl. +apply Props.Dec.F.add_m;[apply Regs.eq_refl|assumption]. +intros;apply RegRegProps.add_add. +intros;apply RegRegProps.Dec.F.add_m;[apply eq_refl|assumption]. + +cut (VertexSet.In (fst_ext e) (V_interf_reg_mreg g) /\ + VertexSet.In (snd_ext e) (V_interf_reg_mreg g)). +intro H1. destruct H1 as [H1 H2]. split; +(apply VertexSet.union_3; apply VertexSet.union_2; assumption). +unfold V_interf_reg_mreg. +unfold IE_reg_mreg in H0. +rewrite SetRegMreg.fold_1 in *. +induction (SetRegMreg.elements (interf_reg_mreg g)). +simpl in H0. elim (EdgeSet.empty_1 H0). +rewrite MEdgeFacts.fold_left_assoc in H0. +destruct (proj1 (RegRegProps.Dec.F.add_iff _ _ _) H0). +fold (eq (Regs.reg_to_Reg (fst a), Regs.mreg_to_Reg (snd a), Some N0) e) in H1. +rewrite RegFacts.fold_left_assoc. +destruct (eq_charac _ _ H1); destruct H2 as [H2 H3]. +change_rewrite. split. +apply VertexSet.add_1. assumption. +apply VertexSet.add_2. apply VertexSet.add_1. assumption. +change_rewrite. split. +apply VertexSet.add_2. apply VertexSet.add_1. assumption. +apply VertexSet.add_1. assumption. +intros y z s. +rewrite (Props.add_add (VertexSet.add (Regs.mreg_to_Reg (snd y)) s) + (Regs.mreg_to_Reg (snd z)) + (Regs.reg_to_Reg (fst y))). +rewrite (Props.add_add s + (Regs.mreg_to_Reg (snd z)) + (Regs.mreg_to_Reg (snd y))). +set (s' := VertexSet.add (Regs.mreg_to_Reg (snd z) ) s). +rewrite (Props.add_add s' + (Regs.mreg_to_Reg (snd y)) + (Regs.reg_to_Reg (fst z))). +apply Props.add_add. +intros. apply Props.Dec.F.add_m. +apply Regs.eq_refl. +apply Props.Dec.F.add_m;[apply Regs.eq_refl|assumption]. +rewrite RegFacts.fold_left_assoc. +split; do 2 apply VertexSet.add_2. +apply (proj1 (IHl H1)). +apply (proj2 (IHl H1)). +intros y z s. +rewrite (Props.add_add (VertexSet.add (Regs.mreg_to_Reg (snd y)) s) + (Regs.mreg_to_Reg (snd z)) + (Regs.reg_to_Reg (fst y))). +rewrite (Props.add_add s + (Regs.mreg_to_Reg (snd z)) + (Regs.mreg_to_Reg (snd y))). +set (s' := VertexSet.add (Regs.mreg_to_Reg (snd z) ) s). +rewrite (Props.add_add s' + (Regs.mreg_to_Reg (snd y)) + (Regs.reg_to_Reg (fst z))). +apply Props.add_add. +intros. apply Props.Dec.F.add_m. +apply Regs.eq_refl. +apply Props.Dec.F.add_m;[apply Regs.eq_refl|assumption]. +intros;apply RegRegProps.add_add. +intros;apply RegRegProps.Dec.F.add_m;[apply eq_refl|assumption]. +Qed. +*) + +Lemma vertices_empty_map : forall x, +VertexMap.In x empty_map <-> +VertexSet.In x interfgraph_vertices. + +Proof. +unfold empty_map. split; intros. +set (f := (fun (x : VertexSet.elt) (m : VertexMap.t VertexSet.t) => + VertexMap.add x VertexSet.empty m)) in *. +rewrite VertexSet.fold_1 in *. +set (f' := fun a e => f e a) in *. +generalize VertexSet.elements_2. intro HH. +generalize (HH interfgraph_vertices). clear HH. intro HH. +induction (VertexSet.elements interfgraph_vertices). simpl in H. +rewrite MapFacts.empty_in_iff in H. inversion H. +cut (EqualSetMap (fold_left f' (a :: l) (VertexMap.empty VertexSet.t)) + (f' (fold_left f' l (VertexMap.empty VertexSet.t)) a)). intro. +generalize (H0 x). clear H0. intro H0. inversion H0. subst. +simpl in H. rewrite MapFacts.in_find_iff in H. rewrite <-H2 in H. +elim H. auto. +set (tmp := fold_left f' l (VertexMap.empty VertexSet.t)) in *. +unfold f', f in H2. +destruct (Vertex.eq_dec x a). +rewrite e. apply HH. left. apply Regs.eq_refl. +apply IHl. +rewrite MapFacts.add_neq_o in H2. +rewrite MapFacts.in_find_iff. rewrite <-H2. congruence. +auto. + +intros. apply HH. right. auto. + +apply fold_left_assoc_map. +intros. +unfold f'. unfold f. +unfold EqualSetMap. intro. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y). +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +intro. elim n. rewrite H0. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Regs.eq_dec x0 y). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s). +constructor. apply VertexSet.eq_refl. +constructor. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f', f. +destruct (Vertex.eq_dec x0 a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. + +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +auto. +auto. +auto. + +set (f := (fun (x : VertexSet.elt) (m : VertexMap.t VertexSet.t) => + VertexMap.add x VertexSet.empty m)) in *. +rewrite VertexSet.fold_1. +set (f' := fun a e => f e a) in *. +generalize VertexSet.elements_1. intro HH. +generalize (HH _ _ H). clear HH. intro HH. +induction (VertexSet.elements interfgraph_vertices). +inversion HH. +cut (EqualSetMap (fold_left f' (a :: l) (VertexMap.empty VertexSet.t)) + (f' (fold_left f' l (VertexMap.empty VertexSet.t)) a)). intro. +generalize (H0 x). clear H0. intro H0. inversion H0. subst. +set (tmp := fold_left f' l (VertexMap.empty VertexSet.t)) in *. +unfold f', f in H3. +destruct (Vertex.eq_dec x a). +simpl. rewrite MapFacts.in_find_iff. +rewrite MapFacts.add_eq_o in H3. congruence. +apply Regs.eq_sym. auto. +inversion HH; subst. +elim (n H4). +rewrite MapFacts.add_neq_o in H3. +rewrite MapFacts.in_find_iff in IHl. rewrite <-H3 in IHl. +elim (IHl H4 (refl_equal _)). +auto. +simpl. rewrite MapFacts.in_find_iff. rewrite <-H1. congruence. + +apply fold_left_assoc_map. +intros. +unfold f'. unfold f. +unfold EqualSetMap. intro. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y). +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +intro. elim n. rewrite H0. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Regs.eq_dec x0 y). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s). +constructor. apply VertexSet.eq_refl. +constructor. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f', f. +destruct (Vertex.eq_dec x0 a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. + +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +auto. +auto. +auto. +Qed. + + +Lemma extremities_imap : forall x, +VertexMap.In x interf_map <-> +VertexSet.In x interfgraph_vertices. + +Proof. +split; intros. +unfold interf_map in H. +unfold set2map in H. +set (f := (fun (e : EdgeSet.elt) (m : VertexMap.t VertexSet.t) => + VertexMap.add (fst_ext e) (new_adj_set (fst_ext e) (snd_ext e) m) + (VertexMap.add (snd_ext e) + (new_adj_set (snd_ext e) (fst_ext e) m) m))) in *. +rewrite EdgeSet.fold_1 in H. +set (f' := fun a e => f e a) in *. +cut (forall e, EdgeSet.In e interfgraph_interference_edges -> + VertexSet.In (fst_ext e) interfgraph_vertices + /\ VertexSet.In (snd_ext e) interfgraph_vertices). intro. + +generalize EdgeSet.elements_2. intro HH. +generalize (HH interfgraph_interference_edges). clear HH. intro HH. + +induction (EdgeSet.elements interfgraph_interference_edges). simpl in H. +apply (proj1 (vertices_empty_map x)). auto. + +cut (EqualSetMap (fold_left f' (a :: l) empty_map) + (f' (fold_left f' l empty_map ) a)). intro. +generalize (H1 x). clear H1. intro H1. inversion H1. subst. +simpl in H. rewrite MapFacts.in_find_iff in H. rewrite <-H3 in H. +elim H. auto. + +set (tmp := fold_left f' l empty_map) in *. +unfold f' in H3. unfold f in H3. +destruct (Vertex.eq_dec x (fst_ext a)). +rewrite e. apply (H0 a). +apply HH. left. apply E.eq_refl. +destruct (Regs.eq_dec x (snd_ext a)). +rewrite e. apply (H0 a). +apply HH. left. apply E.eq_refl. +apply IHl. +do 2 rewrite MapFacts.add_neq_o in H3. +rewrite MapFacts.in_find_iff. rewrite <-H3. congruence. +auto. +auto. +auto. + +intros. apply HH. right. auto. + +apply fold_left_assoc_map. +intros. +unfold f'. unfold f. +unfold EqualSetMap. intro. +destruct (Vertex.eq_dec x0 (fst_ext z)). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 (fst_ext y)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +rewrite (MapFacts.find_o _ (Regs.eq_sym e)). +rewrite (MapFacts.find_o _ (Regs.eq_sym e0)). +destruct (VertexMap.find x0 s). +apply Props.add_add. +do 2 rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. auto. +rewrite <-e0. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (snd_ext y)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +rewrite (MapFacts.find_o _ (Regs.eq_sym e)). +rewrite (MapFacts.find_o _ (Regs.eq_sym e0)). +destruct (VertexMap.find x0 s). +apply Props.add_add. +do 2 rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. auto. +rewrite <-e0. auto. +intro. elim n. rewrite e. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply VertexSet.eq_refl. +intro. elim n0. rewrite e. auto. +intro. elim n. rewrite e. auto. +apply (Regs.eq_sym e). +intro. elim n0. auto. +intro. elim n. auto. +apply (Regs.eq_sym e). +rewrite MapFacts.add_neq_o. +destruct (Regs.eq_dec x0 (snd_ext z)). +rewrite MapFacts.add_eq_o. +destruct (Regs.eq_dec x0 (fst_ext y)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite (MapFacts.find_o _ (Regs.eq_sym e)). +rewrite (MapFacts.find_o _ (Regs.eq_sym e0)). +destruct (VertexMap.find x0 s). +apply Props.add_add. +do 2 rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. auto. +rewrite <-e0. auto. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (snd_ext y)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite (MapFacts.find_o _ (Regs.eq_sym e)). +rewrite (MapFacts.find_o _ (Regs.eq_sym e0)). +destruct (VertexMap.find x0 s). +apply Props.add_add. +do 2 rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. auto. +rewrite <-e0. auto. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +intro. elim n0. rewrite H1. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply VertexSet.eq_refl. +intro. elim n1. rewrite H1. auto. +intro. elim n0. rewrite H1. auto. +apply Regs.eq_sym. auto. +auto. +auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (fst_ext y)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply VertexSet.eq_refl. +intro. elim n0. rewrite H1. auto. +intro. elim n. rewrite H1. auto. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (snd_ext y)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +constructor. +apply VertexSet.eq_refl. +intro. elim n0. rewrite H1. auto. +intro. elim n. rewrite H1. auto. +apply Regs.eq_sym. auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s). +constructor. apply VertexSet.eq_refl. +constructor. +auto. +auto. +auto. +auto. +auto. +auto. +auto. +auto. + +intros. +unfold f'. unfold f. unfold EqualSetMap. intros. +destruct (Regs.eq_dec x0 (fst_ext a0)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +generalize (H1 (fst_ext a0)). intro. +inversion H2. +apply VertexSet.eq_refl. +apply Props.Dec.F.add_m. apply Regs.eq_refl. auto. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Regs.eq_dec x0 (snd_ext a0)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +generalize (H1 (snd_ext a0)). intro. +inversion H2. +apply VertexSet.eq_refl. +apply Props.Dec.F.add_m. apply Regs.eq_refl. auto. +apply Regs.eq_sym. auto. +intro. elim n. rewrite H2. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +auto. +intro. elim n0. rewrite H2. auto. +intro. elim n. rewrite H2. auto. +auto. +auto. + +intros. apply extremities_interf_graph. right. auto. + +rewrite <-vertices_empty_map in H. +unfold interf_map. unfold set2map. +set (f := (fun (e : EdgeSet.elt) (m : VertexMap.t VertexSet.t) => + VertexMap.add (fst_ext e) (new_adj_set (fst_ext e) (snd_ext e) m) + (VertexMap.add (snd_ext e) (new_adj_set (snd_ext e) (fst_ext e) m) m))) in *. +rewrite EdgeSet.fold_1. +set (f' := fun e a => f a e) in *. +induction (EdgeSet.elements interfgraph_interference_edges). simpl. auto. +cut (EqualSetMap (fold_left f' (a :: l) empty_map) + (f' (fold_left f' l empty_map) a)). intro. +generalize (H0 x). clear H0. intro H0. inversion H0. subst. +set (tmp := fold_left f' l empty_map) in *. +unfold f' in H3. unfold f in H3. +destruct (Vertex.eq_dec x (fst_ext a)). +rewrite MapFacts.add_eq_o in H3. congruence. +apply (Regs.eq_sym e). +destruct (Vertex.eq_dec x (snd_ext a)). +rewrite MapFacts.add_neq_o in H3. +rewrite MapFacts.add_eq_o in H3. +congruence. +apply (Regs.eq_sym e). +auto. +rewrite MapFacts.add_neq_o in H3. +rewrite MapFacts.add_neq_o in H3. +rewrite MapFacts.in_find_iff in IHl. rewrite <-H3 in IHl. +congruence. +auto. +auto. + +simpl. rewrite MapFacts.in_find_iff. rewrite <-H1. congruence. +apply fold_left_assoc_map. +intros. +unfold f'. unfold f. +unfold EqualSetMap. intro. +destruct (Vertex.eq_dec x0 (fst_ext z)). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 (fst_ext y)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +rewrite (MapFacts.find_o _ (Regs.eq_sym e)). +rewrite (MapFacts.find_o _ (Regs.eq_sym e0)). +destruct (VertexMap.find x0 s). +apply Props.add_add. +do 2 rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. auto. +rewrite <-e0. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (snd_ext y)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +rewrite (MapFacts.find_o _ (Regs.eq_sym e)). +rewrite (MapFacts.find_o _ (Regs.eq_sym e0)). +destruct (VertexMap.find x0 s). +apply Props.add_add. +do 2 rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. auto. +rewrite <-e0. auto. +intro. elim n. rewrite e. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply VertexSet.eq_refl. +intro. elim n0. rewrite e. auto. +intro. elim n. rewrite e. auto. +apply (Regs.eq_sym e). +intro. elim n0. auto. +intro. elim n. auto. +apply (Regs.eq_sym e). +rewrite MapFacts.add_neq_o. +destruct (Regs.eq_dec x0 (snd_ext z)). +rewrite MapFacts.add_eq_o. +destruct (Regs.eq_dec x0 (fst_ext y)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite (MapFacts.find_o _ (Regs.eq_sym e)). +rewrite (MapFacts.find_o _ (Regs.eq_sym e0)). +destruct (VertexMap.find x0 s). +apply Props.add_add. +do 2 rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. auto. +rewrite <-e0. auto. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (snd_ext y)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite (MapFacts.find_o _ (Regs.eq_sym e)). +rewrite (MapFacts.find_o _ (Regs.eq_sym e0)). +destruct (VertexMap.find x0 s). +apply Props.add_add. +do 2 rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. auto. +rewrite <-e0. auto. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +intro. elim n0. rewrite H0. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply VertexSet.eq_refl. +intro. elim n1. rewrite H0. auto. +intro. elim n0. rewrite H0. auto. +apply Regs.eq_sym. auto. +auto. +auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (fst_ext y)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply VertexSet.eq_refl. +intro. elim n0. rewrite H0. auto. +intro. elim n. rewrite H0. auto. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (snd_ext y)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +constructor. +apply VertexSet.eq_refl. +intro. elim n0. rewrite H0. auto. +intro. elim n. rewrite H0. auto. +apply Regs.eq_sym. auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s). +constructor. apply VertexSet.eq_refl. +constructor. +auto. +auto. +auto. +auto. +auto. +auto. +auto. +auto. + +intros. +unfold f'. unfold f. unfold EqualSetMap. intros. +destruct (Regs.eq_dec x0 (fst_ext a0)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +generalize (H0 (fst_ext a0)). intro. +inversion H1. +apply VertexSet.eq_refl. +apply Props.Dec.F.add_m. apply Regs.eq_refl. auto. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Regs.eq_dec x0 (snd_ext a0)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +generalize (H0 (snd_ext a0)). intro. +inversion H1. +apply VertexSet.eq_refl. +apply Props.Dec.F.add_m. apply Regs.eq_refl. auto. +apply Regs.eq_sym. auto. +intro. elim n. rewrite H1. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +auto. +intro. elim n0. rewrite H1. auto. +intro. elim n. rewrite H1. auto. +auto. +auto. +Qed. + +Lemma extremities_pmap : forall x, +VertexMap.In x pref_map <-> +VertexSet.In x interfgraph_vertices. + +Proof. +split; intros. +unfold pref_map in H. +unfold set2map in H. +set (f := (fun (e : EdgeSet.elt) (m : VertexMap.t VertexSet.t) => + VertexMap.add (fst_ext e) (new_adj_set (fst_ext e) (snd_ext e) m) + (VertexMap.add (snd_ext e) + (new_adj_set (snd_ext e) (fst_ext e) m) m))) in *. +rewrite EdgeSet.fold_1 in H. +set (f' := fun a e => f e a) in *. +cut (forall e, EdgeSet.In e interfgraph_affinity_edges -> + VertexSet.In (fst_ext e) interfgraph_vertices + /\ VertexSet.In (snd_ext e) interfgraph_vertices). intro. + +generalize EdgeSet.elements_2. intro HH. +generalize (HH interfgraph_affinity_edges). clear HH. intro HH. + +induction (EdgeSet.elements interfgraph_affinity_edges). simpl in H. +apply (proj1 (vertices_empty_map x)). auto. + +cut (EqualSetMap (fold_left f' (a :: l) empty_map) + (f' (fold_left f' l empty_map) a)). intro. +generalize (H1 x). clear H1. intro H1. inversion H1. subst. +simpl in H. rewrite MapFacts.in_find_iff in H. rewrite <-H3 in H. +elim H. auto. + +set (tmp := fold_left f' l empty_map) in *. +unfold f' in H3. unfold f in H3. +destruct (Vertex.eq_dec x (fst_ext a)). +rewrite e. apply (H0 a). +apply HH. left. apply E.eq_refl. +destruct (Regs.eq_dec x (snd_ext a)). +rewrite e. apply (H0 a). +apply HH. left. apply E.eq_refl. +apply IHl. +do 2 rewrite MapFacts.add_neq_o in H3. +rewrite MapFacts.in_find_iff. rewrite <-H3. congruence. +auto. +auto. +auto. + +intros. apply HH. right. auto. + +apply fold_left_assoc_map. +intros. +unfold f'. unfold f. +unfold EqualSetMap. intro. +destruct (Vertex.eq_dec x0 (fst_ext z)). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 (fst_ext y)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +rewrite (MapFacts.find_o _ (Regs.eq_sym e)). +rewrite (MapFacts.find_o _ (Regs.eq_sym e0)). +destruct (VertexMap.find x0 s). +apply Props.add_add. +do 2 rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. auto. +rewrite <-e0. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (snd_ext y)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +rewrite (MapFacts.find_o _ (Regs.eq_sym e)). +rewrite (MapFacts.find_o _ (Regs.eq_sym e0)). +destruct (VertexMap.find x0 s). +apply Props.add_add. +do 2 rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. auto. +rewrite <-e0. auto. +intro. elim n. rewrite e. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply VertexSet.eq_refl. +intro. elim n0. rewrite e. auto. +intro. elim n. rewrite e. auto. +apply (Regs.eq_sym e). +intro. elim n0. auto. +intro. elim n. auto. +apply (Regs.eq_sym e). +rewrite MapFacts.add_neq_o. +destruct (Regs.eq_dec x0 (snd_ext z)). +rewrite MapFacts.add_eq_o. +destruct (Regs.eq_dec x0 (fst_ext y)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite (MapFacts.find_o _ (Regs.eq_sym e)). +rewrite (MapFacts.find_o _ (Regs.eq_sym e0)). +destruct (VertexMap.find x0 s). +apply Props.add_add. +do 2 rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. auto. +rewrite <-e0. auto. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (snd_ext y)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite (MapFacts.find_o _ (Regs.eq_sym e)). +rewrite (MapFacts.find_o _ (Regs.eq_sym e0)). +destruct (VertexMap.find x0 s). +apply Props.add_add. +do 2 rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. auto. +rewrite <-e0. auto. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +intro. elim n0. rewrite H1. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply VertexSet.eq_refl. +intro. elim n1. rewrite H1. auto. +intro. elim n0. rewrite H1. auto. +apply Regs.eq_sym. auto. +auto. +auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (fst_ext y)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply VertexSet.eq_refl. +intro. elim n0. rewrite H1. auto. +intro. elim n. rewrite H1. auto. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (snd_ext y)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +constructor. +apply VertexSet.eq_refl. +intro. elim n0. rewrite H1. auto. +intro. elim n. rewrite H1. auto. +apply Regs.eq_sym. auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s). +constructor. apply VertexSet.eq_refl. +constructor. +auto. +auto. +auto. +auto. +auto. +auto. +auto. +auto. + +intros. +unfold f'. unfold f. unfold EqualSetMap. intros. +destruct (Regs.eq_dec x0 (fst_ext a0)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +generalize (H1 (fst_ext a0)). intro. +inversion H2. +apply VertexSet.eq_refl. +apply Props.Dec.F.add_m. apply Regs.eq_refl. auto. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Regs.eq_dec x0 (snd_ext a0)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +generalize (H1 (snd_ext a0)). intro. +inversion H2. +apply VertexSet.eq_refl. +apply Props.Dec.F.add_m. apply Regs.eq_refl. auto. +apply Regs.eq_sym. auto. +intro. elim n. rewrite H2. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +auto. +intro. elim n0. rewrite H2. auto. +intro. elim n. rewrite H2. auto. +auto. +auto. + +intros. apply extremities_interf_graph. left. auto. + +rewrite <-vertices_empty_map in H. +unfold pref_map. unfold set2map. +set (f := (fun (e : EdgeSet.elt) (m : VertexMap.t VertexSet.t) => + VertexMap.add (fst_ext e) (new_adj_set (fst_ext e) (snd_ext e) m) + (VertexMap.add (snd_ext e) (new_adj_set (snd_ext e) (fst_ext e) m) m))) in *. +rewrite EdgeSet.fold_1. +set (f' := fun e a => f a e) in *. +induction (EdgeSet.elements interfgraph_affinity_edges). simpl. auto. +cut (EqualSetMap (fold_left f' (a :: l) empty_map) + (f' (fold_left f' l empty_map) a)). intro. +generalize (H0 x). clear H0. intro H0. inversion H0. subst. +set (tmp := fold_left f' l empty_map) in *. +unfold f' in H3. unfold f in H3. +destruct (Vertex.eq_dec x (fst_ext a)). +rewrite MapFacts.add_eq_o in H3. congruence. +apply (Regs.eq_sym e). +destruct (Vertex.eq_dec x (snd_ext a)). +rewrite MapFacts.add_neq_o in H3. +rewrite MapFacts.add_eq_o in H3. +congruence. +apply (Regs.eq_sym e). +auto. +rewrite MapFacts.add_neq_o in H3. +rewrite MapFacts.add_neq_o in H3. +rewrite MapFacts.in_find_iff in IHl. rewrite <-H3 in IHl. +congruence. +auto. +auto. + +simpl. rewrite MapFacts.in_find_iff. rewrite <-H1. congruence. +apply fold_left_assoc_map. +intros. +unfold f'. unfold f. +unfold EqualSetMap. intro. +destruct (Vertex.eq_dec x0 (fst_ext z)). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 (fst_ext y)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +rewrite (MapFacts.find_o _ (Regs.eq_sym e)). +rewrite (MapFacts.find_o _ (Regs.eq_sym e0)). +destruct (VertexMap.find x0 s). +apply Props.add_add. +do 2 rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. auto. +rewrite <-e0. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (snd_ext y)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +rewrite (MapFacts.find_o _ (Regs.eq_sym e)). +rewrite (MapFacts.find_o _ (Regs.eq_sym e0)). +destruct (VertexMap.find x0 s). +apply Props.add_add. +do 2 rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. auto. +rewrite <-e0. auto. +intro. elim n. rewrite e. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply VertexSet.eq_refl. +intro. elim n0. rewrite e. auto. +intro. elim n. rewrite e. auto. +apply (Regs.eq_sym e). +intro. elim n0. auto. +intro. elim n. auto. +apply (Regs.eq_sym e). +rewrite MapFacts.add_neq_o. +destruct (Regs.eq_dec x0 (snd_ext z)). +rewrite MapFacts.add_eq_o. +destruct (Regs.eq_dec x0 (fst_ext y)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite (MapFacts.find_o _ (Regs.eq_sym e)). +rewrite (MapFacts.find_o _ (Regs.eq_sym e0)). +destruct (VertexMap.find x0 s). +apply Props.add_add. +do 2 rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. auto. +rewrite <-e0. auto. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (snd_ext y)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite (MapFacts.find_o _ (Regs.eq_sym e)). +rewrite (MapFacts.find_o _ (Regs.eq_sym e0)). +destruct (VertexMap.find x0 s). +apply Props.add_add. +do 2 rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. auto. +rewrite <-e0. auto. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +intro. elim n0. rewrite H0. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply VertexSet.eq_refl. +intro. elim n1. rewrite H0. auto. +intro. elim n0. rewrite H0. auto. +apply Regs.eq_sym. auto. +auto. +auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (fst_ext y)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply VertexSet.eq_refl. +intro. elim n0. rewrite H0. auto. +intro. elim n. rewrite H0. auto. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (snd_ext y)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +constructor. +apply VertexSet.eq_refl. +intro. elim n0. rewrite H0. auto. +intro. elim n. rewrite H0. auto. +apply Regs.eq_sym. auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s). +constructor. apply VertexSet.eq_refl. +constructor. +auto. +auto. +auto. +auto. +auto. +auto. +auto. +auto. + +intros. +unfold f'. unfold f. unfold EqualSetMap. intros. +destruct (Regs.eq_dec x0 (fst_ext a0)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +generalize (H0 (fst_ext a0)). intro. +inversion H1. +apply VertexSet.eq_refl. +apply Props.Dec.F.add_m. apply Regs.eq_refl. auto. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Regs.eq_dec x0 (snd_ext a0)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +generalize (H0 (snd_ext a0)). intro. +inversion H1. +apply VertexSet.eq_refl. +apply Props.Dec.F.add_m. apply Regs.eq_refl. auto. +apply Regs.eq_sym. auto. +intro. elim n. rewrite H1. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +auto. +intro. elim n0. rewrite H1. auto. +intro. elim n. rewrite H1. auto. +auto. +auto. +Qed. + +Lemma set2map_charac : forall x y s m, +VertexSet.In x (adj_set y (set2map s m)) -> +(exists w, EdgeSet.In (x,y,w) s) \/ VertexSet.In x (adj_set y m). + +Proof. +intros. +unfold set2map in H. +set (f := (fun (e : EdgeSet.elt) (m : VertexMap.t VertexSet.t) => + VertexMap.add (fst_ext e) + (new_adj_set (fst_ext e) (snd_ext e) m) + (VertexMap.add (snd_ext e) + (new_adj_set (snd_ext e) (fst_ext e) m) m))) in *. +rewrite EdgeSet.fold_1 in H. +set (f' := fun a e => f e a) in *. +generalize EdgeSet.elements_2. intro HH. +generalize (HH s). clear HH. intro HH. +induction (EdgeSet.elements s). simpl in H. right. auto. +cut (EqualSetMap (fold_left f' (a :: l) m) + (f' (fold_left f' l m) a)). intro. +generalize (H0 y). clear H0. intro H0. inversion H0. subst. +set (tmp := fold_left f' l m) in *. +unfold f', f in H3. +destruct (Vertex.eq_dec y (fst_ext a)). +rewrite MapFacts.add_eq_o in H3. +congruence. +apply Regs.eq_sym. auto. +destruct (Vertex.eq_dec y (snd_ext a)). +rewrite MapFacts.add_neq_o in H3. +rewrite MapFacts.add_eq_o in H3. +congruence. +apply Regs.eq_sym. auto. +auto. +rewrite MapFacts.add_neq_o in H3. +rewrite MapFacts.add_neq_o in H3. +simpl in H. unfold adj_set in H. rewrite <-H2 in H. +elim (VertexSet.empty_1 H). +auto. +auto. +simpl in H. unfold adj_set in H. rewrite <-H1 in H. +rewrite H3 in H. +set (tmp := fold_left f' l m) in *. +unfold f', f in H2. +destruct (Vertex.eq_dec y (fst_ext a)). +rewrite MapFacts.add_eq_o in H2. inversion H2. +rewrite H5 in H. +unfold new_adj_set in H. +case_eq (VertexMap.find (fst_ext a) tmp); intros. +rewrite H4 in H. +destruct (Regs.eq_dec (snd_ext a) x). +left. exists (get_weight a). +apply HH. left. +fold (eq (x,y,get_weight a) a). +rewrite edge_comm. Eq_eq. +generalize (VertexSet.add_3 n H). intro. +apply IHl. +unfold adj_set. rewrite (MapFacts.find_o _ e). +rewrite H4. auto. +intros. apply HH. right. auto. +rewrite H4 in H. +left. exists (get_weight a). +apply HH. left. +fold (eq (x,y,get_weight a) a). +rewrite edge_comm. Eq_eq. +apply Regs.eq_sym. apply VertexSet.singleton_1. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o in H2. +destruct (Vertex.eq_dec y (snd_ext a)). +rewrite MapFacts.add_eq_o in H2. +inversion H2. subst. clear H2. +unfold new_adj_set in H. +case_eq (VertexMap.find (snd_ext a) tmp); intros. +rewrite H2 in H. +destruct (Vertex.eq_dec (fst_ext a) x). +left. exists (get_weight a). +apply HH. left. +fold (eq (x,y,get_weight a) a). Eq_eq. +apply IHl. +generalize (VertexSet.add_3 n0 H). intro. +unfold new_adj_set in H3. rewrite H2 in H3. +unfold adj_set. rewrite (MapFacts.find_o _ e). rewrite H2. +auto. +intros. apply HH. right. auto. +rewrite H2 in H. +left. exists (get_weight a). +apply HH. left. +fold (eq (x,y,get_weight a) a). Eq_eq. +apply Regs.eq_sym. apply VertexSet.singleton_1. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o in H2. +apply IHl. +unfold adj_set. rewrite <-H2. auto. +intros. apply HH. right. auto. +auto. +auto. + +apply fold_left_assoc_map. +unfold f', f, EqualSetMap. intros. +destruct (Vertex.eq_dec x0 (fst_ext z)). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 (fst_ext y0)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +rewrite <-(MapFacts.find_o _ e). +rewrite <-(MapFacts.find_o _ e0). +destruct (VertexMap.find x0 s0). +apply Props.add_add. +rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (snd_ext y0)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +rewrite <-(MapFacts.find_o _ e). +rewrite <-(MapFacts.find_o _ e0). +destruct (VertexMap.find x0 s0). +apply Props.add_add. +rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +intro. elim n. rewrite H0. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply VertexSet.eq_refl. +intro. elim n0. rewrite H0. auto. +intro. elim n. rewrite H0. auto. +apply Regs.eq_sym. auto. +auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (snd_ext z)). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 (fst_ext y0)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite <-(MapFacts.find_o _ e). +rewrite <-(MapFacts.find_o _ e0). +destruct (VertexMap.find x0 s0). apply Props.add_add. +rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +intro. elim n. rewrite H0. auto. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (snd_ext y0)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite <-(MapFacts.find_o _ e). +rewrite <-(MapFacts.find_o _ e0). +destruct (VertexMap.find x0 s0). +apply Props.add_add. +rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +intro. elim n. rewrite H0. auto. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +intro. elim n0. rewrite H0. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply VertexSet.eq_refl. +intro. elim n1. rewrite H0. auto. +intro. elim n0. rewrite H0. auto. +apply Regs.eq_sym. auto. +auto. +auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (fst_ext y0)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply VertexSet.eq_refl. +intro. elim n0. rewrite H0. auto. +intro. elim n. rewrite H0. auto. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (snd_ext y0)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply VertexSet.eq_refl. +intro. elim n0. rewrite H0. auto. +intro. elim n. rewrite H0. auto. +apply Regs.eq_sym. auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s0); constructor. +apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. +auto. +auto. +auto. +auto. + +intros. +unfold EqualSetMap, f', f. intros. +destruct (Vertex.eq_dec x0 (fst_ext a0)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +generalize (H0 (fst_ext a0)). intro. +inversion H1. apply VertexSet.eq_refl. +apply Props.Dec.F.add_m. +apply Regs.eq_refl. +auto. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (snd_ext a0)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +generalize (H0 (snd_ext a0)). intro. +inversion H1. +apply VertexSet.eq_refl. +apply Props.Dec.F.add_m. +apply Regs.eq_refl. +auto. +apply Regs.eq_sym. auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +auto. +auto. +auto. +auto. +auto. +Qed. + +Lemma adj_set_empty : forall x, +VertexSet.Equal (adj_set x empty_map) VertexSet.empty. + +Proof. +intros. +unfold empty_map. +set (f := (fun (x0 : VertexSet.elt) (m : VertexMap.t VertexSet.t) => + VertexMap.add x0 VertexSet.empty m)) in *. +rewrite VertexSet.fold_1. +set (f' := fun a e => f e a) in *. +induction (VertexSet.elements interfgraph_vertices). +simpl. unfold adj_set. simpl. reflexivity. +cut (EqualSetMap (fold_left f' (a :: l) (VertexMap.empty VertexSet.t)) + (f' (fold_left f' l (VertexMap.empty VertexSet.t)) a)). intro. +generalize (H x). clear H. intro H. inversion H. subst. +set (tmp := fold_left f' l (VertexMap.empty VertexSet.t)) in *. +unfold f', f in H2. +destruct (Regs.eq_dec x a). rewrite MapFacts.add_eq_o in H2. congruence. +apply Regs.eq_sym. auto. +simpl. unfold adj_set. rewrite <-H1. apply VertexSet.eq_refl. +set (tmp := fold_left f' l (VertexMap.empty VertexSet.t)) in *. +unfold f', f in H1. +destruct (Regs.eq_dec x a). rewrite MapFacts.add_eq_o in H1. +unfold adj_set. simpl. rewrite <-H0. inversion H1. +rewrite H4 in H2. auto. +apply Regs.eq_sym. auto. +simpl. unfold adj_set. rewrite <-H0. +rewrite MapFacts.add_neq_o in H1. +unfold adj_set in IHl. rewrite <-H1 in IHl. +rewrite H2. auto. +auto. + +apply fold_left_assoc_map. +intros. +unfold f'. unfold f. +unfold EqualSetMap. intro. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y). +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +intro. elim n. rewrite H. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Regs.eq_dec x0 y). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s). +constructor. apply VertexSet.eq_refl. +constructor. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f', f. +destruct (Vertex.eq_dec x0 a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. + +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +auto. +auto. +auto. +Qed. + +Lemma affinity_weights_interf_graph : +forall e, EdgeSet.In e interfgraph_affinity_edges -> +get_weight e = Some N0. + +Proof. +intros e H. +unfold interfgraph_affinity_edges in H. +generalize (proj1 (resolve_conflicts_2 _ _ _ H)). clear H. intro H. +destruct (EdgeSet.union_1 H). +unfold AE_reg_reg in H0. +rewrite SetRegReg.fold_1 in H0. +induction (SetRegReg.elements (pref_reg_reg g)). +simpl in H0. elim (EdgeSet.empty_1 H0). +rewrite MEdgeFacts.fold_left_assoc in H0. +destruct (proj1 (RegRegProps.Dec.F.add_iff _ _ _) H0). +fold (eq (Regs.reg_to_Reg (fst a), Regs.reg_to_Reg (snd a), Some N0) e) in H1. +rewrite <-(get_weight_m _ _ H1). auto. +apply IHl. assumption. +intros;apply RegRegProps.add_add. +intros;apply RegRegProps.Dec.F.add_m;[apply eq_refl|assumption]. +unfold AE_reg_mreg in H0. +rewrite SetRegMreg.fold_1 in H0. +induction (SetRegMreg.elements (pref_reg_mreg g)). +simpl in H0. elim (EdgeSet.empty_1 H0). +rewrite MEdgeFacts.fold_left_assoc in H0. +destruct (proj1 (RegRegProps.Dec.F.add_iff _ _ _) H0). +fold (eq (Regs.reg_to_Reg (fst a), Regs.mreg_to_Reg (snd a), Some N0) e) in H1. +rewrite <-(get_weight_m _ _ H1). auto. +apply IHl. assumption. +intros;apply RegRegProps.add_add. +intros;apply RegRegProps.Dec.F.add_m;[apply eq_refl|assumption]. +Qed. + +Lemma interference_weights_interf_graph : +forall e, EdgeSet.In e interfgraph_interference_edges -> +get_weight e = None. + +Proof. +intros e H. +unfold interfgraph_affinity_edges in H. +destruct (EdgeSet.union_1 H). +unfold IE_reg_reg in H0. +rewrite SetRegReg.fold_1 in H0. +induction (SetRegReg.elements (interf_reg_reg g)). +simpl in H0. elim (EdgeSet.empty_1 H0). +rewrite MEdgeFacts.fold_left_assoc in H0. +destruct (proj1 (RegRegProps.Dec.F.add_iff _ _ _) H0). +fold (eq (Regs.reg_to_Reg (fst a), Regs.reg_to_Reg (snd a), Some N0) e) in H1. +rewrite <-(get_weight_m _ _ H1). auto. +apply IHl. assumption. +intros;apply RegRegProps.add_add. +intros;apply RegRegProps.Dec.F.add_m;[apply eq_refl|assumption]. +unfold IE_reg_mreg in H0. +rewrite SetRegMreg.fold_1 in H0. +induction (SetRegMreg.elements (interf_reg_mreg g)). +simpl in H0. elim (EdgeSet.empty_1 H0). +rewrite MEdgeFacts.fold_left_assoc in H0. +destruct (proj1 (RegRegProps.Dec.F.add_iff _ _ _) H0). +fold (eq (Regs.reg_to_Reg (fst a), Regs.mreg_to_Reg (snd a), Some N0) e) in H1. +rewrite <-(get_weight_m _ _ H1). auto. +apply IHl. assumption. +intros;apply RegRegProps.add_add. +intros;apply RegRegProps.Dec.F.add_m;[apply eq_refl|assumption]. +Qed. + +Lemma simple_graph : forall x y, + VertexSet.In x (adj_set y interf_map ) /\ + VertexSet.In x (adj_set y pref_map ) -> False. + +Proof. +intros. destruct H. +unfold interf_map in H. +generalize (set2map_charac x y interfgraph_interference_edges empty_map H). +clear H. intro H. destruct H. destruct H. +unfold pref_map in H0. +generalize (set2map_charac x y interfgraph_affinity_edges empty_map H0). +clear H0. intro H0. destruct H0. destruct H0. +generalize (resolve_conflicts_2 (x,y,x1) _ _ H0). intro. +destruct H1. elim H2. +change_rewrite. +generalize (interference_weights_interf_graph (x,y,x0) H). simpl. intro. +rewrite H3 in H. assumption. +rewrite adj_set_empty in H0. elim (VertexSet.empty_1 H0). +rewrite adj_set_empty in H. elim (VertexSet.empty_1 H). +Qed. + +Lemma sym_imap : forall x y, + VertexSet.In x (adj_set y interf_map ) -> + VertexSet.In y (adj_set x interf_map). + +Proof. +intros. +unfold interf_map, set2map in *. +set (f := (fun (e : EdgeSet.elt) (m : VertexMap.t VertexSet.t) => + VertexMap.add (fst_ext e) (new_adj_set (fst_ext e) (snd_ext e) m) + (VertexMap.add (snd_ext e) (new_adj_set (snd_ext e) (fst_ext e) m) + m))) in *. +rewrite EdgeSet.fold_1. rewrite EdgeSet.fold_1 in H. +set (f' := fun a e => f e a) in *. +induction (EdgeSet.elements interfgraph_interference_edges). +simpl in H. +rewrite adj_set_empty in H. elim (VertexSet.empty_1 H). +cut (EqualSetMap (fold_left f' (a :: l) empty_map) + (f' (fold_left f' l empty_map) a)). intro. +generalize (H0 y). intro. inversion H1. simpl in *. +unfold adj_set in H. rewrite <-H3 in H. +elim (VertexSet.empty_1 H). +set (tmp := fold_left f' l empty_map) in *. +unfold f', f in H3. +destruct (Vertex.eq_dec y (fst_ext a)). +rewrite MapFacts.add_eq_o in H3. +unfold adj_set in H. simpl in H. rewrite <-H2 in H. +inversion H3. subst. clear H3. +rewrite H4 in H. +unfold new_adj_set in H. +generalize (H0 x). intro. inversion H3. subst. +unfold f', f in H7. +destruct (Vertex.eq_dec x (fst_ext a)). +rewrite MapFacts.add_eq_o in H7. congruence. +apply Regs.eq_sym. auto. +destruct (Vertex.eq_dec (snd_ext a) x). +rewrite MapFacts.add_neq_o in H7. rewrite MapFacts.add_eq_o in H7. +congruence. +apply Regs.eq_sym. auto. +auto. +rewrite MapFacts.add_neq_o in H7. +rewrite MapFacts.add_neq_o in H7. +case_eq (VertexMap.find (fst_ext a) tmp); intros; rewrite H5 in H. +generalize (VertexSet.add_3 n0 H). clear H. intro H. +assert (VertexSet.In y (adj_set x tmp)). +apply IHl. unfold adj_set. rewrite (MapFacts.find_o _ e). +rewrite H5. auto. +unfold adj_set in H8. rewrite <-H7 in H8. elim (VertexSet.empty_1 H8). +elim (n0 (VertexSet.singleton_1 H)). +auto. +auto. + +simpl. unfold adj_set. rewrite <-H5. +rewrite H7. +unfold f', f in H6. +destruct (Vertex.eq_dec x (fst_ext a)). +rewrite MapFacts.add_eq_o in H6. +inversion H6. subst. clear H6. +unfold new_adj_set. case_eq (VertexMap.find (fst_ext a) tmp); intros. +rewrite H6 in H. +destruct (Vertex.eq_dec y (snd_ext a)). +apply VertexSet.add_1. apply Regs.eq_sym. auto. +apply VertexSet.add_2. +assert (VertexSet.In y (adj_set x tmp)). +apply IHl. unfold adj_set. +rewrite (MapFacts.find_o _ e). +rewrite H6. +destruct (Vertex.eq_dec (snd_ext a) x). +elim n. rewrite e. rewrite <-e0. apply Regs.eq_sym. auto. +apply (VertexSet.add_3 n0 H). +unfold adj_set in H8. +rewrite (MapFacts.find_o _ e0) in H8. rewrite H6 in H8. auto. +rewrite H6 in H. +rewrite e. rewrite <-e0. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o in H6. +destruct (Vertex.eq_dec (snd_ext a) x). +rewrite MapFacts.add_eq_o in H6. +inversion H6. subst. clear H6. +unfold new_adj_set. +destruct (VertexMap.find (snd_ext a) tmp). +apply VertexSet.add_1. +apply Regs.eq_sym. auto. +apply VertexSet.singleton_2. apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o in H6. +clear H0 H1. +assert (VertexSet.In y (adj_set x tmp)). +apply IHl. +unfold adj_set. rewrite (MapFacts.find_o _ e). +case_eq (VertexMap.find (fst_ext a) tmp); intros. +rewrite H0 in H. +apply (VertexSet.add_3 n0 H). +rewrite H0 in H. elim (n0 (VertexSet.singleton_1 H)). +unfold adj_set in H0. rewrite <-H6 in H0. auto. +auto. +auto. +apply Regs.eq_sym. auto. + +rewrite MapFacts.add_neq_o in H3. +destruct (Vertex.eq_dec y (snd_ext a)). +rewrite MapFacts.add_eq_o in H3. +unfold adj_set in H. simpl in H. rewrite <-H2 in H. +rewrite H4 in H. +inversion H3. subst. clear H3. +unfold new_adj_set in H. +case_eq (VertexMap.find (snd_ext a) tmp); intros. +rewrite H3 in H. +destruct (Vertex.eq_dec (fst_ext a) x). +unfold adj_set. simpl. +generalize (H0 x). clear H0. intro. inversion H0. +unfold f',f in H7. +rewrite MapFacts.add_eq_o in H7. +congruence. +apply Regs.eq_sym. auto. +unfold f', f in H6. +rewrite MapFacts.add_eq_o in H6. inversion H6. subst. clear H6. +rewrite H7. +unfold new_adj_set in H4. rewrite H3 in H4. +unfold new_adj_set. +destruct (VertexMap.find (fst_ext a) tmp). +apply VertexSet.add_1. apply Regs.eq_sym. auto. +apply VertexSet.singleton_2. apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +unfold adj_set. simpl. +generalize (H0 x). clear H0. intro. inversion H0. +unfold f',f in H7. +rewrite MapFacts.add_neq_o in H7. +destruct (Vertex.eq_dec (snd_ext a) x). +rewrite MapFacts.add_eq_o in H7. congruence. auto. +rewrite MapFacts.add_neq_o in H7. +unfold adj_set in IHl. rewrite <-H7 in IHl. apply IHl. +rewrite (MapFacts.find_o _ e). rewrite H3. +apply (VertexSet.add_3 n0 H). +auto. +auto. +unfold f', f in H6. +rewrite MapFacts.add_neq_o in H6. +destruct (Vertex.eq_dec x (snd_ext a)). +rewrite MapFacts.add_eq_o in H6. inversion H6. subst. clear H6. +rewrite H7. +unfold new_adj_set. rewrite H3. +apply VertexSet.add_2. +unfold adj_set in IHl. rewrite <-(MapFacts.find_o _ e0) in H3. +rewrite H3 in IHl. apply IHl. +rewrite (MapFacts.find_o _ e). rewrite <-(MapFacts.find_o _ e0). +rewrite H3. +apply (VertexSet.add_3 n0 H). +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o in H6. +rewrite H7. +unfold adj_set in IHl. rewrite <-H6 in IHl. +apply IHl. rewrite (MapFacts.find_o _ e). +rewrite H3. +apply (VertexSet.add_3 n0 H). +auto. +auto. +rewrite H3 in H. +generalize (H0 x). intro. inversion H5. subst. +unfold f',f in H8. +rewrite MapFacts.add_eq_o in H8. +congruence. +apply VertexSet.singleton_1. auto. +unfold f', f in H7. +rewrite MapFacts.add_eq_o in H7. +inversion H7. subst. clear H7. +simpl. unfold adj_set. rewrite <-H6. +rewrite H8. unfold new_adj_set. +destruct (VertexMap.find (fst_ext a) tmp). +apply VertexSet.add_1. apply Regs.eq_sym. auto. +apply VertexSet.singleton_2. apply Regs.eq_sym. auto. +apply (VertexSet.singleton_1 H). +apply Regs.eq_sym. auto. + +rewrite MapFacts.add_neq_o in H3. +generalize (H0 x). intro. inversion H5. subst. +unfold f', f in H8. +destruct (Vertex.eq_dec x (fst_ext a)). +rewrite MapFacts.add_eq_o in H8. +congruence. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o in H8. +destruct (Vertex.eq_dec x (snd_ext a)). +rewrite MapFacts.add_eq_o in H8. +congruence. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o in H8. +simpl in H. unfold adj_set in H. rewrite <-H2 in H. +unfold adj_set in IHl. +assert (VertexSet.In y VertexSet.empty). +rewrite <-H8 in IHl. apply IHl. +rewrite <-H3. rewrite <-H4. auto. +elim (VertexSet.empty_1 H6). +auto. +auto. +unfold f',f in H7. +destruct (Vertex.eq_dec (fst_ext a) x). +rewrite MapFacts.add_eq_o in H7. +inversion H7. subst. clear H7. +unfold adj_set. simpl. rewrite <-H6. +rewrite H8. +unfold new_adj_set. +simpl in H. unfold adj_set in H. rewrite <-H2 in H. +case_eq (VertexMap.find (fst_ext a) tmp); intros. +unfold adj_set in IHl. +rewrite <-(MapFacts.find_o _ e) in IHl. +rewrite H7 in IHl. +apply VertexSet.add_2. apply IHl. +rewrite <-H3. rewrite <-H4. auto. +assert (VertexSet.In y VertexSet.empty). +unfold adj_set in IHl. rewrite (MapFacts.find_o _ e) in H7. +rewrite H7 in IHl. apply IHl. +rewrite <-H3. rewrite <-H4. auto. +elim (VertexSet.empty_1 H9). +auto. +rewrite MapFacts.add_neq_o in H7. +destruct (Vertex.eq_dec (snd_ext a) x). +rewrite MapFacts.add_eq_o in H7. +inversion H7. subst. clear H7. +simpl. unfold adj_set. rewrite <-H6. +rewrite H8. unfold new_adj_set. +simpl in H. unfold adj_set in H. rewrite <-H2 in H. +rewrite (MapFacts.find_o _ e). +unfold adj_set in IHl. +case_eq (VertexMap.find x tmp); intros. +rewrite H7 in IHl. +apply VertexSet.add_2. +apply IHl. +rewrite <-H3. rewrite <-H4. auto. +assert (VertexSet.In y VertexSet.empty). +unfold adj_set in IHl. rewrite H7 in IHl. +apply IHl. rewrite <-H3. rewrite <-H4. auto. +elim (VertexSet.empty_1 H9). +auto. +rewrite MapFacts.add_neq_o in H7. +simpl. unfold adj_set. rewrite <-H6. +rewrite H8. +unfold adj_set in IHl. rewrite <-H7 in IHl. rewrite <-H3 in IHl. +apply IHl. rewrite <-H4. simpl in H. unfold adj_set in H. rewrite <-H2 in H. auto. +auto. +auto. +auto. +auto. + +apply fold_left_assoc_map. +intros. +unfold f'. unfold f. +unfold EqualSetMap. intro. +destruct (Vertex.eq_dec x0 (fst_ext z)). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 (fst_ext y0)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +rewrite (MapFacts.find_o _ (Regs.eq_sym e)). +rewrite (MapFacts.find_o _ (Regs.eq_sym e0)). +destruct (VertexMap.find x0 s). +apply Props.add_add. +do 2 rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. auto. +rewrite <-e0. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (snd_ext y0)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +rewrite (MapFacts.find_o _ (Regs.eq_sym e)). +rewrite (MapFacts.find_o _ (Regs.eq_sym e0)). +destruct (VertexMap.find x0 s). +apply Props.add_add. +do 2 rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. auto. +rewrite <-e0. auto. +intro. elim n. rewrite e. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply VertexSet.eq_refl. +intro. elim n0. rewrite e. auto. +intro. elim n. rewrite e. auto. +apply (Regs.eq_sym e). +intro. elim n0. auto. +intro. elim n. auto. +apply (Regs.eq_sym e). +rewrite MapFacts.add_neq_o. +destruct (Regs.eq_dec x0 (snd_ext z)). +rewrite MapFacts.add_eq_o. +destruct (Regs.eq_dec x0 (fst_ext y0)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite (MapFacts.find_o _ (Regs.eq_sym e)). +rewrite (MapFacts.find_o _ (Regs.eq_sym e0)). +destruct (VertexMap.find x0 s). +apply Props.add_add. +do 2 rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. auto. +rewrite <-e0. auto. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (snd_ext y0)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite (MapFacts.find_o _ (Regs.eq_sym e)). +rewrite (MapFacts.find_o _ (Regs.eq_sym e0)). +destruct (VertexMap.find x0 s). +apply Props.add_add. +do 2 rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. auto. +rewrite <-e0. auto. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +intro. elim n0. rewrite H0. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply VertexSet.eq_refl. +intro. elim n1. rewrite H0. auto. +intro. elim n0. rewrite H0. auto. +apply Regs.eq_sym. auto. +auto. +auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (fst_ext y0)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply VertexSet.eq_refl. +intro. elim n0. rewrite H0. auto. +intro. elim n. rewrite H0. auto. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (snd_ext y0)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +constructor. +apply VertexSet.eq_refl. +intro. elim n0. rewrite H0. auto. +intro. elim n. rewrite H0. auto. +apply Regs.eq_sym. auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s). +constructor. apply VertexSet.eq_refl. +constructor. +auto. +auto. +auto. +auto. +auto. +auto. +auto. +auto. + +intros. +unfold f'. unfold f. unfold EqualSetMap. intros. +destruct (Regs.eq_dec x0 (fst_ext a0)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +generalize (H0 (fst_ext a0)). intro. +inversion H1. +apply VertexSet.eq_refl. +apply Props.Dec.F.add_m. apply Regs.eq_refl. auto. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Regs.eq_dec x0 (snd_ext a0)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +generalize (H0 (snd_ext a0)). intro. +inversion H1. +apply VertexSet.eq_refl. +apply Props.Dec.F.add_m. apply Regs.eq_refl. auto. +apply Regs.eq_sym. auto. +intro. elim n. rewrite H1. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +auto. +intro. elim n0. rewrite H1. auto. +intro. elim n. rewrite H1. auto. +auto. +auto. +Qed. + +Lemma sym_pmap : forall x y, + VertexSet.In x (adj_set y pref_map) -> + VertexSet.In y (adj_set x pref_map). + +Proof. +intros. +unfold pref_map, set2map in *. +set (f := (fun (e : EdgeSet.elt) (m : VertexMap.t VertexSet.t) => + VertexMap.add (fst_ext e) (new_adj_set (fst_ext e) (snd_ext e) m) + (VertexMap.add (snd_ext e) (new_adj_set (snd_ext e) (fst_ext e) m) + m))) in *. +rewrite EdgeSet.fold_1. rewrite EdgeSet.fold_1 in H. +set (f' := fun a e => f e a) in *. +induction (EdgeSet.elements interfgraph_affinity_edges). +simpl in H. +rewrite adj_set_empty in H. elim (VertexSet.empty_1 H). +cut (EqualSetMap (fold_left f' (a :: l) empty_map) + (f' (fold_left f' l empty_map) a)). intro. +generalize (H0 y). intro. inversion H1. simpl in *. +unfold adj_set in H. rewrite <-H3 in H. +elim (VertexSet.empty_1 H). +set (tmp := fold_left f' l empty_map) in *. +unfold f', f in H3. +destruct (Vertex.eq_dec y (fst_ext a)). +rewrite MapFacts.add_eq_o in H3. +unfold adj_set in H. simpl in H. rewrite <-H2 in H. +inversion H3. subst. clear H3. +rewrite H4 in H. +unfold new_adj_set in H. +generalize (H0 x). intro. inversion H3. subst. +unfold f', f in H7. +destruct (Vertex.eq_dec x (fst_ext a)). +rewrite MapFacts.add_eq_o in H7. congruence. +apply Regs.eq_sym. auto. +destruct (Vertex.eq_dec (snd_ext a) x). +rewrite MapFacts.add_neq_o in H7. rewrite MapFacts.add_eq_o in H7. +congruence. +apply Regs.eq_sym. auto. +auto. +rewrite MapFacts.add_neq_o in H7. +rewrite MapFacts.add_neq_o in H7. +case_eq (VertexMap.find (fst_ext a) tmp); intros; rewrite H5 in H. +generalize (VertexSet.add_3 n0 H). clear H. intro H. +assert (VertexSet.In y (adj_set x tmp)). +apply IHl. unfold adj_set. rewrite (MapFacts.find_o _ e). +rewrite H5. auto. +unfold adj_set in H8. rewrite <-H7 in H8. elim (VertexSet.empty_1 H8). +elim (n0 (VertexSet.singleton_1 H)). +auto. +auto. + +simpl. unfold adj_set. rewrite <-H5. +rewrite H7. +unfold f', f in H6. +destruct (Vertex.eq_dec x (fst_ext a)). +rewrite MapFacts.add_eq_o in H6. +inversion H6. subst. clear H6. +unfold new_adj_set. case_eq (VertexMap.find (fst_ext a) tmp); intros. +rewrite H6 in H. +destruct (Vertex.eq_dec y (snd_ext a)). +apply VertexSet.add_1. apply Regs.eq_sym. auto. +apply VertexSet.add_2. +assert (VertexSet.In y (adj_set x tmp)). +apply IHl. unfold adj_set. +rewrite (MapFacts.find_o _ e). +rewrite H6. +destruct (Vertex.eq_dec (snd_ext a) x). +elim n. rewrite e. rewrite <-e0. apply Regs.eq_sym. auto. +apply (VertexSet.add_3 n0 H). +unfold adj_set in H8. +rewrite (MapFacts.find_o _ e0) in H8. rewrite H6 in H8. auto. +rewrite H6 in H. +rewrite e. rewrite <-e0. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o in H6. +destruct (Vertex.eq_dec (snd_ext a) x). +rewrite MapFacts.add_eq_o in H6. +inversion H6. subst. clear H6. +unfold new_adj_set. +destruct (VertexMap.find (snd_ext a) tmp). +apply VertexSet.add_1. +apply Regs.eq_sym. auto. +apply VertexSet.singleton_2. apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o in H6. +clear H0 H1. +assert (VertexSet.In y (adj_set x tmp)). +apply IHl. +unfold adj_set. rewrite (MapFacts.find_o _ e). +case_eq (VertexMap.find (fst_ext a) tmp); intros. +rewrite H0 in H. +apply (VertexSet.add_3 n0 H). +rewrite H0 in H. elim (n0 (VertexSet.singleton_1 H)). +unfold adj_set in H0. rewrite <-H6 in H0. auto. +auto. +auto. +apply Regs.eq_sym. auto. + +rewrite MapFacts.add_neq_o in H3. +destruct (Vertex.eq_dec y (snd_ext a)). +rewrite MapFacts.add_eq_o in H3. +unfold adj_set in H. simpl in H. rewrite <-H2 in H. +rewrite H4 in H. +inversion H3. subst. clear H3. +unfold new_adj_set in H. +case_eq (VertexMap.find (snd_ext a) tmp); intros. +rewrite H3 in H. +destruct (Vertex.eq_dec (fst_ext a) x). +unfold adj_set. simpl. +generalize (H0 x). clear H0. intro. inversion H0. +unfold f',f in H7. +rewrite MapFacts.add_eq_o in H7. +congruence. +apply Regs.eq_sym. auto. +unfold f', f in H6. +rewrite MapFacts.add_eq_o in H6. inversion H6. subst. clear H6. +rewrite H7. +unfold new_adj_set in H4. rewrite H3 in H4. +unfold new_adj_set. +destruct (VertexMap.find (fst_ext a) tmp). +apply VertexSet.add_1. apply Regs.eq_sym. auto. +apply VertexSet.singleton_2. apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +unfold adj_set. simpl. +generalize (H0 x). clear H0. intro. inversion H0. +unfold f',f in H7. +rewrite MapFacts.add_neq_o in H7. +destruct (Vertex.eq_dec (snd_ext a) x). +rewrite MapFacts.add_eq_o in H7. congruence. auto. +rewrite MapFacts.add_neq_o in H7. +unfold adj_set in IHl. rewrite <-H7 in IHl. apply IHl. +rewrite (MapFacts.find_o _ e). rewrite H3. +apply (VertexSet.add_3 n0 H). +auto. +auto. +unfold f', f in H6. +rewrite MapFacts.add_neq_o in H6. +destruct (Vertex.eq_dec x (snd_ext a)). +rewrite MapFacts.add_eq_o in H6. inversion H6. subst. clear H6. +rewrite H7. +unfold new_adj_set. rewrite H3. +apply VertexSet.add_2. +unfold adj_set in IHl. rewrite <-(MapFacts.find_o _ e0) in H3. +rewrite H3 in IHl. apply IHl. +rewrite (MapFacts.find_o _ e). rewrite <-(MapFacts.find_o _ e0). +rewrite H3. +apply (VertexSet.add_3 n0 H). +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o in H6. +rewrite H7. +unfold adj_set in IHl. rewrite <-H6 in IHl. +apply IHl. rewrite (MapFacts.find_o _ e). +rewrite H3. +apply (VertexSet.add_3 n0 H). +auto. +auto. +rewrite H3 in H. +generalize (H0 x). intro. inversion H5. subst. +unfold f',f in H8. +rewrite MapFacts.add_eq_o in H8. +congruence. +apply VertexSet.singleton_1. auto. +unfold f', f in H7. +rewrite MapFacts.add_eq_o in H7. +inversion H7. subst. clear H7. +simpl. unfold adj_set. rewrite <-H6. +rewrite H8. unfold new_adj_set. +destruct (VertexMap.find (fst_ext a) tmp). +apply VertexSet.add_1. apply Regs.eq_sym. auto. +apply VertexSet.singleton_2. apply Regs.eq_sym. auto. +apply (VertexSet.singleton_1 H). +apply Regs.eq_sym. auto. + +rewrite MapFacts.add_neq_o in H3. +generalize (H0 x). intro. inversion H5. subst. +unfold f', f in H8. +destruct (Vertex.eq_dec x (fst_ext a)). +rewrite MapFacts.add_eq_o in H8. +congruence. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o in H8. +destruct (Vertex.eq_dec x (snd_ext a)). +rewrite MapFacts.add_eq_o in H8. +congruence. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o in H8. +simpl in H. unfold adj_set in H. rewrite <-H2 in H. +unfold adj_set in IHl. +assert (VertexSet.In y VertexSet.empty). +rewrite <-H8 in IHl. apply IHl. +rewrite <-H3. rewrite <-H4. auto. +elim (VertexSet.empty_1 H6). +auto. +auto. +unfold f',f in H7. +destruct (Vertex.eq_dec (fst_ext a) x). +rewrite MapFacts.add_eq_o in H7. +inversion H7. subst. clear H7. +unfold adj_set. simpl. rewrite <-H6. +rewrite H8. +unfold new_adj_set. +simpl in H. unfold adj_set in H. rewrite <-H2 in H. +case_eq (VertexMap.find (fst_ext a) tmp); intros. +unfold adj_set in IHl. +rewrite <-(MapFacts.find_o _ e) in IHl. +rewrite H7 in IHl. +apply VertexSet.add_2. apply IHl. +rewrite <-H3. rewrite <-H4. auto. +assert (VertexSet.In y VertexSet.empty). +unfold adj_set in IHl. rewrite (MapFacts.find_o _ e) in H7. +rewrite H7 in IHl. apply IHl. +rewrite <-H3. rewrite <-H4. auto. +elim (VertexSet.empty_1 H9). +auto. +rewrite MapFacts.add_neq_o in H7. +destruct (Vertex.eq_dec (snd_ext a) x). +rewrite MapFacts.add_eq_o in H7. +inversion H7. subst. clear H7. +simpl. unfold adj_set. rewrite <-H6. +rewrite H8. unfold new_adj_set. +simpl in H. unfold adj_set in H. rewrite <-H2 in H. +rewrite (MapFacts.find_o _ e). +unfold adj_set in IHl. +case_eq (VertexMap.find x tmp); intros. +rewrite H7 in IHl. +apply VertexSet.add_2. +apply IHl. +rewrite <-H3. rewrite <-H4. auto. +assert (VertexSet.In y VertexSet.empty). +unfold adj_set in IHl. rewrite H7 in IHl. +apply IHl. rewrite <-H3. rewrite <-H4. auto. +elim (VertexSet.empty_1 H9). +auto. +rewrite MapFacts.add_neq_o in H7. +simpl. unfold adj_set. rewrite <-H6. +rewrite H8. +unfold adj_set in IHl. rewrite <-H7 in IHl. rewrite <-H3 in IHl. +apply IHl. rewrite <-H4. simpl in H. unfold adj_set in H. rewrite <-H2 in H. auto. +auto. +auto. +auto. +auto. + +apply fold_left_assoc_map. +intros. +unfold f'. unfold f. +unfold EqualSetMap. intro. +destruct (Vertex.eq_dec x0 (fst_ext z)). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 (fst_ext y0)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +rewrite (MapFacts.find_o _ (Regs.eq_sym e)). +rewrite (MapFacts.find_o _ (Regs.eq_sym e0)). +destruct (VertexMap.find x0 s). +apply Props.add_add. +do 2 rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. auto. +rewrite <-e0. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (snd_ext y0)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +rewrite (MapFacts.find_o _ (Regs.eq_sym e)). +rewrite (MapFacts.find_o _ (Regs.eq_sym e0)). +destruct (VertexMap.find x0 s). +apply Props.add_add. +do 2 rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. auto. +rewrite <-e0. auto. +intro. elim n. rewrite e. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply VertexSet.eq_refl. +intro. elim n0. rewrite e. auto. +intro. elim n. rewrite e. auto. +apply (Regs.eq_sym e). +intro. elim n0. auto. +intro. elim n. auto. +apply (Regs.eq_sym e). +rewrite MapFacts.add_neq_o. +destruct (Regs.eq_dec x0 (snd_ext z)). +rewrite MapFacts.add_eq_o. +destruct (Regs.eq_dec x0 (fst_ext y0)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite (MapFacts.find_o _ (Regs.eq_sym e)). +rewrite (MapFacts.find_o _ (Regs.eq_sym e0)). +destruct (VertexMap.find x0 s). +apply Props.add_add. +do 2 rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. auto. +rewrite <-e0. auto. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (snd_ext y0)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite (MapFacts.find_o _ (Regs.eq_sym e)). +rewrite (MapFacts.find_o _ (Regs.eq_sym e0)). +destruct (VertexMap.find x0 s). +apply Props.add_add. +do 2 rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. auto. +rewrite <-e0. auto. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +intro. elim n0. rewrite H0. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply VertexSet.eq_refl. +intro. elim n1. rewrite H0. auto. +intro. elim n0. rewrite H0. auto. +apply Regs.eq_sym. auto. +auto. +auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (fst_ext y0)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply VertexSet.eq_refl. +intro. elim n0. rewrite H0. auto. +intro. elim n. rewrite H0. auto. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (snd_ext y0)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +constructor. +apply VertexSet.eq_refl. +intro. elim n0. rewrite H0. auto. +intro. elim n. rewrite H0. auto. +apply Regs.eq_sym. auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s). +constructor. apply VertexSet.eq_refl. +constructor. +auto. +auto. +auto. +auto. +auto. +auto. +auto. +auto. + +intros. +unfold f'. unfold f. unfold EqualSetMap. intros. +destruct (Regs.eq_dec x0 (fst_ext a0)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +generalize (H0 (fst_ext a0)). intro. +inversion H1. +apply VertexSet.eq_refl. +apply Props.Dec.F.add_m. apply Regs.eq_refl. auto. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Regs.eq_dec x0 (snd_ext a0)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +generalize (H0 (snd_ext a0)). intro. +inversion H1. +apply VertexSet.eq_refl. +apply Props.Dec.F.add_m. apply Regs.eq_refl. auto. +apply Regs.eq_sym. auto. +intro. elim n. rewrite H1. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +auto. +intro. elim n0. rewrite H1. auto. +intro. elim n. rewrite H1. auto. +auto. +auto. +Qed. + +Lemma set_reg_reg_diff_ext : forall x g, +SetRegReg.In x (interf_reg_reg g) \/ +SetRegReg.In x (pref_reg_reg g) -> fst x <> snd x. + +Proof. +Admitted. + +Lemma IE_reg_reg_diff : forall x y w, +EdgeSet.In (x,y,w) IE_reg_reg -> ~Vertex.eq x y. + +Proof. +intros. +unfold IE_reg_reg in H. +set (f := (fun (e : SetRegReg.elt) (s : EdgeSet.t) => + EdgeSet.add + (Regs.reg_to_Reg (fst e), Regs.reg_to_Reg (snd e), None) s)) in *. +rewrite SetRegReg.fold_1 in H. +set (f' := fun a e => f e a) in *. +generalize SetRegReg.elements_2. intro HH. +generalize (HH (interf_reg_reg g)). clear HH. intro HH. +induction (SetRegReg.elements (interf_reg_reg g)). simpl in H. +elim (EdgeSet.empty_1 H). +rewrite MEdgeFacts.fold_left_assoc in H. +set (tmp := fold_left f' l EdgeSet.empty) in *. +unfold f', f in H. +destruct (proj1 (RegRegProps.Dec.F.add_iff _ _ _) H). +fold (eq (Regs.reg_to_Reg (fst a), Regs.reg_to_Reg (snd a), None) (x,y,w)) in H0. +destruct (eq_charac _ _ H0); change_rewrite; destruct H1. +assert (fst a <> snd a). +apply set_reg_reg_diff_ext with (g:=g). +left. apply HH. +left. auto. +intro. rewrite <-H1 in H4. rewrite <-H2 in H4. +inversion H4. subst. elim (H3 H7). +assert (fst a <> snd a). +apply set_reg_reg_diff_ext with (g:=g). +left. apply HH. +left. auto. +intro. rewrite <-H1 in H4. rewrite <-H2 in H4. +inversion H4. subst. elim H3. auto. +apply IHl. auto. + +intros. apply HH. right. auto. + +unfold f', f. intros. +apply RegRegProps.add_add. + +unfold f', f. intros. +apply RegRegProps.Dec.F.add_m. apply E.eq_refl. auto. +Qed. + +Lemma IE_reg_mreg_diff : forall x y w, +EdgeSet.In (x,y,w) IE_reg_mreg -> ~Vertex.eq x y. + +Proof. +intros. +unfold IE_reg_mreg in H. +set (f := (fun (e : SetRegMreg.elt) (s : EdgeSet.t) => + EdgeSet.add + (Regs.reg_to_Reg (fst e), Regs.mreg_to_Reg (snd e), None) s)) in *. +rewrite SetRegMreg.fold_1 in H. +set (f' := fun a e => f e a) in *. +generalize SetRegMreg.elements_2. intro HH. +generalize (HH (interf_reg_mreg g)). clear HH. intro HH. +induction (SetRegMreg.elements (interf_reg_mreg g)). simpl in H. +elim (EdgeSet.empty_1 H). +rewrite MEdgeFacts.fold_left_assoc in H. +set (tmp := fold_left f' l EdgeSet.empty) in *. +unfold f', f in H. +destruct (proj1 (RegRegProps.Dec.F.add_iff _ _ _) H). +fold (eq (Regs.reg_to_Reg (fst a), Regs.mreg_to_Reg (snd a), None) (x,y,w)) in H0. +destruct (eq_charac _ _ H0); change_rewrite; destruct H1. +intro. rewrite <-H1 in H3. rewrite <-H2 in H3. inversion H3. +intro. rewrite <-H1 in H3. rewrite <-H2 in H3. inversion H3. +apply IHl. auto. + +intros. apply HH. right. auto. + +unfold f', f. intros. +apply RegRegProps.add_add. + +unfold f', f. intros. +apply RegRegProps.Dec.F.add_m. apply E.eq_refl. auto. +Qed. + +Lemma AE_reg_reg_diff : forall x y w, +EdgeSet.In (x,y,w) AE_reg_reg -> ~Vertex.eq x y. + +Proof. +intros. +unfold AE_reg_reg in H. +set (f := (fun (e : SetRegReg.elt) (s : EdgeSet.t) => + EdgeSet.add + (Regs.reg_to_Reg (fst e), Regs.reg_to_Reg (snd e), Some N0) s)) in *. +rewrite SetRegReg.fold_1 in H. +set (f' := fun a e => f e a) in *. +generalize SetRegReg.elements_2. intro HH. +generalize (HH (pref_reg_reg g)). clear HH. intro HH. +induction (SetRegReg.elements (pref_reg_reg g)). simpl in H. +elim (EdgeSet.empty_1 H). +rewrite MEdgeFacts.fold_left_assoc in H. +set (tmp := fold_left f' l EdgeSet.empty) in *. +unfold f', f in H. +destruct (proj1 (RegRegProps.Dec.F.add_iff _ _ _) H). +fold (eq (Regs.reg_to_Reg (fst a), Regs.reg_to_Reg (snd a), None) (x,y,w)) in H0. +destruct (eq_charac _ _ H0); change_rewrite; destruct H1. +assert (fst a <> snd a). +apply set_reg_reg_diff_ext with (g:=g). +right. apply HH. +left. auto. +intro. rewrite <-H1 in H4. rewrite <-H2 in H4. +inversion H4. subst. elim (H3 H7). +assert (fst a <> snd a). +apply set_reg_reg_diff_ext with (g:=g). +right. apply HH. +left. auto. +intro. rewrite <-H1 in H4. rewrite <-H2 in H4. +inversion H4. subst. elim H3. auto. +apply IHl. auto. + +intros. apply HH. right. auto. + +unfold f', f. intros. +apply RegRegProps.add_add. + +unfold f', f. intros. +apply RegRegProps.Dec.F.add_m. apply E.eq_refl. auto. +Qed. + +Lemma AE_reg_mreg_diff : forall x y w, +EdgeSet.In (x,y,w) AE_reg_mreg -> ~Vertex.eq x y. + +Proof. +intros. +unfold AE_reg_mreg in H. +set (f := (fun (e : SetRegMreg.elt) (s : EdgeSet.t) => + EdgeSet.add + (Regs.reg_to_Reg (fst e), Regs.mreg_to_Reg (snd e), Some N0) s)) in *. +rewrite SetRegMreg.fold_1 in H. +set (f' := fun a e => f e a) in *. +generalize SetRegMreg.elements_2. intro HH. +generalize (HH (pref_reg_mreg g)). clear HH. intro HH. +induction (SetRegMreg.elements (pref_reg_mreg g)). simpl in H. +elim (EdgeSet.empty_1 H). +rewrite MEdgeFacts.fold_left_assoc in H. +set (tmp := fold_left f' l EdgeSet.empty) in *. +unfold f', f in H. +destruct (proj1 (RegRegProps.Dec.F.add_iff _ _ _) H). +fold (eq (Regs.reg_to_Reg (fst a), Regs.mreg_to_Reg (snd a), None) (x,y,w)) in H0. +destruct (eq_charac _ _ H0); change_rewrite; destruct H1. +intro. rewrite <-H1 in H3. rewrite <-H2 in H3. inversion H3. +intro. rewrite <-H1 in H3. rewrite <-H2 in H3. inversion H3. +apply IHl. auto. + +intros. apply HH. right. auto. + +unfold f', f. intros. +apply RegRegProps.add_add. + +unfold f', f. intros. +apply RegRegProps.Dec.F.add_m. apply E.eq_refl. auto. +Qed. + +Lemma not_eq_extremities : forall x y, +VertexSet.In x (adj_set y interf_map) \/ VertexSet.In x (adj_set y pref_map) -> +~Vertex.eq x y. + +Proof. +intros. +destruct H. +generalize (set2map_charac _ _ _ _ H). intro. +destruct H0. destruct H0. +unfold interfgraph_interference_edges in H0. +destruct (EdgeSet.union_1 H0). +apply (IE_reg_reg_diff x y x0 H1). +apply (IE_reg_mreg_diff x y x0 H1). +rewrite adj_set_empty in H0. elim (VertexSet.empty_1 H0). +generalize (set2map_charac _ _ _ _ H). intro. +destruct H0. destruct H0. +unfold interfgraph_affinity_edges in H0. +generalize (resolve_conflicts_2 _ _ _ H0). intro. +clear H0. destruct H1 as [H0 _]. +destruct (EdgeSet.union_1 H0). +apply (AE_reg_reg_diff x y x0 H1). +apply (AE_reg_mreg_diff x y x0 H1). +rewrite adj_set_empty in H0. elim (VertexSet.empty_1 H0). +Qed. + +Definition graph_translation : new_graph := +let vert := interfgraph_vertices in +let iedges := interfgraph_interference_edges in +let aedges := resolve_conflicts (EdgeSet.union AE_reg_reg AE_reg_mreg) iedges in +let emp := VertexSet.fold + (fun x m => VertexMap.add x VertexSet.empty m) + vert + (VertexMap.empty VertexSet.t) in +let im := set2map iedges emp in +let pm := set2map aedges emp in +Make_Graph + vert im pm + extremities_imap + extremities_pmap + simple_graph + sym_imap + sym_pmap + not_eq_extremities. + +Lemma set2map_charac_2 : forall x y s m, +(exists w, EdgeSet.In (x,y,w) s) \/ VertexSet.In x (adj_set y m) -> +VertexSet.In x (adj_set y (set2map s m)). + +Proof. +intros. +unfold set2map. +set (f := (fun (e : EdgeSet.elt) (m : VertexMap.t VertexSet.t) => + VertexMap.add (fst_ext e) + (new_adj_set (fst_ext e) (snd_ext e) m) + (VertexMap.add (snd_ext e) + (new_adj_set (snd_ext e) (fst_ext e) m) m))) in *. +rewrite EdgeSet.fold_1. +set (f' := fun a e => f e a) in *. +destruct H. destruct H. +generalize EdgeSet.elements_1. intro HH. +generalize (HH s (x,y,x0) H). clear HH. intro HH. +induction (EdgeSet.elements s). +inversion HH. +cut (EqualSetMap (fold_left f' (a :: l) m) + (f' (fold_left f' l m) a)). intro. +generalize (H0 y). clear H0. intro H0. inversion H0. subst. +set (tmp := fold_left f' l m) in *. +unfold f', f in H3. +destruct (Vertex.eq_dec y (fst_ext a)). +rewrite MapFacts.add_eq_o in H3. +congruence. +apply Regs.eq_sym. auto. +destruct (Vertex.eq_dec y (snd_ext a)). +rewrite MapFacts.add_neq_o in H3. +rewrite MapFacts.add_eq_o in H3. +congruence. +apply Regs.eq_sym. auto. +auto. +rewrite MapFacts.add_neq_o in H3. +rewrite MapFacts.add_neq_o in H3. +inversion HH; subst. +destruct (eq_charac _ _ H4); change_rewrite; destruct H4; destruct H1. +elim (n0 H6). +elim (n H6). +unfold adj_set in IHl. rewrite <-H3 in IHl. +elim (VertexSet.empty_1 (IHl H4)). +auto. +auto. + +set (tmp := fold_left f' l m) in *. +unfold adj_set. simpl. rewrite <-H1. rewrite H3. +unfold f',f in H2. +inversion HH; subst. +destruct (eq_charac _ _ H5); change_rewrite; destruct H4. +destruct (Vertex.eq_dec y (fst_ext a)). +rewrite MapFacts.add_eq_o in H2. +inversion H2. subst. clear H2. +unfold new_adj_set. +destruct (VertexMap.find (fst_ext a) tmp). +apply VertexSet.add_1. rewrite H4. rewrite <-H6. rewrite <-e. apply Regs.eq_refl. +apply VertexSet.singleton_2. rewrite H4. rewrite <-H6. rewrite <-e. apply Regs.eq_refl. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o in H2. +rewrite MapFacts.add_eq_o in H2. +inversion H2. subst. clear H2. +unfold new_adj_set. destruct (VertexMap.find (snd_ext a) tmp). +apply VertexSet.add_1. apply Regs.eq_sym. auto. +apply VertexSet.singleton_2. apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +auto. +rewrite MapFacts.add_eq_o in H2. +inversion H2. subst. clear H2. +unfold new_adj_set. destruct (VertexMap.find (fst_ext a) tmp). +apply VertexSet.add_1. apply Regs.eq_sym. auto. +apply VertexSet.singleton_2. apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +destruct (Vertex.eq_dec y (fst_ext a)). +rewrite MapFacts.add_eq_o in H2. +inversion H2. subst. clear H2. +unfold new_adj_set. +rewrite <-(MapFacts.find_o _ e). +case_eq (VertexMap.find y tmp); intros. +apply VertexSet.add_2. +unfold adj_set in IHl. rewrite H2 in IHl. +apply IHl. auto. +assert (VertexSet.In x VertexSet.empty). +unfold adj_set in IHl. rewrite H2 in IHl. apply IHl. auto. +elim (VertexSet.empty_1 H4). +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o in H2. +destruct (Vertex.eq_dec y (snd_ext a)). +rewrite MapFacts.add_eq_o in H2. +inversion H2. subst. clear H2. +unfold new_adj_set. +rewrite <-(MapFacts.find_o _ e). +case_eq (VertexMap.find y tmp); intros. +apply VertexSet.add_2. +unfold adj_set in IHl. rewrite H2 in IHl. apply IHl. auto. +assert (VertexSet.In x VertexSet.empty). +unfold adj_set in IHl. rewrite H2 in IHl. apply IHl. auto. +elim (VertexSet.empty_1 H4). +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o in H2. +unfold adj_set in IHl. rewrite <-H2 in IHl. apply IHl. auto. +auto. +auto. + +apply fold_left_assoc_map. +unfold f', f, EqualSetMap. intros. +destruct (Vertex.eq_dec x1 (fst_ext z)). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x1 (fst_ext y0)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +rewrite <-(MapFacts.find_o _ e). +rewrite <-(MapFacts.find_o _ e0). +destruct (VertexMap.find x1 s0). +apply Props.add_add. +rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x1 (snd_ext y0)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +rewrite <-(MapFacts.find_o _ e). +rewrite <-(MapFacts.find_o _ e0). +destruct (VertexMap.find x1 s0). +apply Props.add_add. +rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +intro. elim n. rewrite H0. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply VertexSet.eq_refl. +intro. elim n0. rewrite H0. auto. +intro. elim n. rewrite H0. auto. +apply Regs.eq_sym. auto. +auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x1 (snd_ext z)). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x1 (fst_ext y0)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite <-(MapFacts.find_o _ e). +rewrite <-(MapFacts.find_o _ e0). +destruct (VertexMap.find x1 s0). apply Props.add_add. +rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +intro. elim n. rewrite H0. auto. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x1 (snd_ext y0)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite <-(MapFacts.find_o _ e). +rewrite <-(MapFacts.find_o _ e0). +destruct (VertexMap.find x1 s0). +apply Props.add_add. +rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +intro. elim n. rewrite H0. auto. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +intro. elim n0. rewrite H0. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply VertexSet.eq_refl. +intro. elim n1. rewrite H0. auto. +intro. elim n0. rewrite H0. auto. +apply Regs.eq_sym. auto. +auto. +auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x1 (fst_ext y0)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply VertexSet.eq_refl. +intro. elim n0. rewrite H0. auto. +intro. elim n. rewrite H0. auto. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x1 (snd_ext y0)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply VertexSet.eq_refl. +intro. elim n0. rewrite H0. auto. +intro. elim n. rewrite H0. auto. +apply Regs.eq_sym. auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x1 s0); constructor. +apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. +auto. +auto. +auto. +auto. + +intros. +unfold EqualSetMap, f', f. intros. +destruct (Vertex.eq_dec x1 (fst_ext a0)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +generalize (H0 (fst_ext a0)). intro. +inversion H1. apply VertexSet.eq_refl. +apply Props.Dec.F.add_m. +apply Regs.eq_refl. +auto. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x1 (snd_ext a0)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +generalize (H0 (snd_ext a0)). intro. +inversion H1. +apply VertexSet.eq_refl. +apply Props.Dec.F.add_m. +apply Regs.eq_refl. +auto. +apply Regs.eq_sym. auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +auto. +auto. +auto. +auto. +auto. + +induction (EdgeSet.elements s). simpl. auto. +cut (EqualSetMap (fold_left f' (a :: l) m) + (f' (fold_left f' l m) a)). intro. +generalize (H0 y). clear H0. intro H0. inversion H0. subst. +set (tmp := fold_left f' l m) in *. +unfold f', f in H3. +destruct (Vertex.eq_dec y (fst_ext a)). +rewrite MapFacts.add_eq_o in H3. +congruence. +apply Regs.eq_sym. auto. +destruct (Vertex.eq_dec y (snd_ext a)). +rewrite MapFacts.add_neq_o in H3. +rewrite MapFacts.add_eq_o in H3. +congruence. +apply Regs.eq_sym. auto. +auto. +rewrite MapFacts.add_neq_o in H3. +rewrite MapFacts.add_neq_o in H3. +unfold adj_set in IHl. rewrite <-H3 in IHl. elim (VertexSet.empty_1 IHl). +auto. +auto. + +unfold adj_set. simpl. rewrite <-H1. +rewrite H3. +set (tmp := fold_left f' l m) in *. +unfold f', f in H2. +destruct (Vertex.eq_dec y (fst_ext a)). +rewrite MapFacts.add_eq_o in H2. +inversion H2. subst. clear H2. +unfold new_adj_set. +rewrite <-(MapFacts.find_o _ e). +case_eq (VertexMap.find y tmp); intros. +apply VertexSet.add_2. +unfold adj_set in IHl. rewrite H2 in IHl. apply IHl. +unfold adj_set in IHl. rewrite H2 in IHl. inversion IHl. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o in H2. +destruct (Vertex.eq_dec y (snd_ext a)). +rewrite MapFacts.add_eq_o in H2. +inversion H2. subst. clear H2. +unfold new_adj_set. +rewrite <-(MapFacts.find_o _ e). +case_eq (VertexMap.find y tmp); intros. +apply VertexSet.add_2. +unfold adj_set in IHl. rewrite H2 in IHl. apply IHl. +unfold adj_set in IHl. rewrite H2 in IHl. inversion IHl. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o in H2. +unfold adj_set in IHl. rewrite <-H2 in IHl. apply IHl. +auto. +auto. + +apply fold_left_assoc_map. +unfold f', f, EqualSetMap. intros. +destruct (Vertex.eq_dec x0 (fst_ext z)). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 (fst_ext y0)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +rewrite <-(MapFacts.find_o _ e). +rewrite <-(MapFacts.find_o _ e0). +destruct (VertexMap.find x0 s0). +apply Props.add_add. +rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (snd_ext y0)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +rewrite <-(MapFacts.find_o _ e). +rewrite <-(MapFacts.find_o _ e0). +destruct (VertexMap.find x0 s0). +apply Props.add_add. +rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +intro. elim n. rewrite H0. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply VertexSet.eq_refl. +intro. elim n0. rewrite H0. auto. +intro. elim n. rewrite H0. auto. +apply Regs.eq_sym. auto. +auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (snd_ext z)). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 (fst_ext y0)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite <-(MapFacts.find_o _ e). +rewrite <-(MapFacts.find_o _ e0). +destruct (VertexMap.find x0 s0). apply Props.add_add. +rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +intro. elim n. rewrite H0. auto. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (snd_ext y0)). +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +rewrite <-(MapFacts.find_o _ e). +rewrite <-(MapFacts.find_o _ e0). +destruct (VertexMap.find x0 s0). +apply Props.add_add. +rewrite Props.singleton_equal_add. apply Props.add_add. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +intro. elim n. rewrite H0. auto. +rewrite <-e. rewrite <-e0. apply Regs.eq_refl. +intro. elim n0. rewrite H0. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply VertexSet.eq_refl. +intro. elim n1. rewrite H0. auto. +intro. elim n0. rewrite H0. auto. +apply Regs.eq_sym. auto. +auto. +auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (fst_ext y0)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply VertexSet.eq_refl. +intro. elim n0. rewrite H0. auto. +intro. elim n. rewrite H0. auto. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (snd_ext y0)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply VertexSet.eq_refl. +intro. elim n0. rewrite H0. auto. +intro. elim n. rewrite H0. auto. +apply Regs.eq_sym. auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s0); constructor. +apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. +auto. +auto. +auto. +auto. + +intros. +unfold EqualSetMap, f', f. intros. +destruct (Vertex.eq_dec x0 (fst_ext a0)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +generalize (H0 (fst_ext a0)). intro. +inversion H1. apply VertexSet.eq_refl. +apply Props.Dec.F.add_m. +apply Regs.eq_refl. +auto. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 (snd_ext a0)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +unfold new_adj_set. +generalize (H0 (snd_ext a0)). intro. +inversion H1. +apply VertexSet.eq_refl. +apply Props.Dec.F.add_m. +apply Regs.eq_refl. +auto. +apply Regs.eq_sym. auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +auto. +auto. +auto. +auto. +auto. +Qed. + +Lemma regreg_IE : forall x y, +SetRegReg.In (x,y) (interf_reg_reg g) -> +EdgeSet.In (Regs.reg_to_Reg x,Regs.reg_to_Reg y,None) IE_reg_reg. + +Proof. +intros. +unfold IE_reg_reg. +set (f := (fun (e : SetRegReg.elt) (s : EdgeSet.t) => + EdgeSet.add (Regs.reg_to_Reg (fst e), Regs.reg_to_Reg (snd e), None) s)) in *. +rewrite SetRegReg.fold_1. +set (f' := fun a e => f e a) in *. +generalize SetRegReg.elements_1. intro HH. +generalize (HH (interf_reg_reg g) (x,y) H). clear HH. intro HH. +induction (SetRegReg.elements (interf_reg_reg g)). simpl. +inversion HH. +rewrite MEdgeFacts.fold_left_assoc. +set (tmp := fold_left f' l EdgeSet.empty) in *. +unfold f', f. +inversion HH; subst. +apply EdgeSet.add_1. destruct H1. simpl in *. subst. apply Edge.eq_refl. +apply EdgeSet.add_2. apply IHl. auto. +unfold f', f. intros. +apply RegRegProps.add_add. + +unfold f', f. intros. +apply RegRegProps.Dec.F.add_m. apply E.eq_refl. auto. +Qed. + +Lemma regmreg_IE : forall x y, +SetRegMreg.In (x,y) (interf_reg_mreg g) -> +EdgeSet.In (Regs.reg_to_Reg x,Regs.mreg_to_Reg y,None) IE_reg_mreg. + +Proof. +intros. +unfold IE_reg_mreg. +set (f := (fun (e : SetRegMreg.elt) (s : EdgeSet.t) => + EdgeSet.add (Regs.reg_to_Reg (fst e), Regs.mreg_to_Reg (snd e), None) s)) in *. +rewrite SetRegMreg.fold_1. +set (f' := fun a e => f e a) in *. +generalize SetRegMreg.elements_1. intro HH. +generalize (HH (interf_reg_mreg g) (x,y) H). clear HH. intro HH. +induction (SetRegMreg.elements (interf_reg_mreg g)). simpl. +inversion HH. +rewrite MEdgeFacts.fold_left_assoc. +set (tmp := fold_left f' l EdgeSet.empty) in *. +unfold f', f. +inversion HH; subst. +apply EdgeSet.add_1. destruct H1. simpl in *. subst. apply Edge.eq_refl. +apply EdgeSet.add_2. apply IHl. auto. +unfold f', f. intros. +apply RegRegProps.add_add. + +unfold f', f. intros. +apply RegRegProps.Dec.F.add_m. apply E.eq_refl. auto. +Qed. + +Lemma regreg_IE_translation : forall x y, +SetRegReg.In (x,y) (interf_reg_reg g) -> +EdgeSet.In (Regs.P x,Regs.P y,None) (IE graph_translation). + +Proof. +intros. +unfold graph_translation. unfold IE. simpl. +rewrite (edgemap_to_edgeset_charac _ _ _ _ sym_imap). +apply set2map_charac_2. +left. exists None. +unfold interfgraph_interference_edges. apply EdgeSet.union_2. +rewrite edge_comm. apply regreg_IE. auto. +Qed. + +Lemma regmreg_IE_translation : forall x y, +SetRegMreg.In (x,y) (interf_reg_mreg g) -> +EdgeSet.In (Regs.P x,Regs.M y,None) (IE graph_translation). + +Proof. +intros. +unfold graph_translation. unfold IE. simpl. +rewrite (edgemap_to_edgeset_charac _ _ _ _ sym_imap). +apply set2map_charac_2. +left. exists None. +unfold interfgraph_interference_edges. apply EdgeSet.union_3. +rewrite edge_comm. apply regmreg_IE. auto. +Qed. + +End Translation. \ No newline at end of file diff --git a/backend/IRC.v b/backend/IRC.v new file mode 100755 index 00000000..7842a4bd --- /dev/null +++ b/backend/IRC.v @@ -0,0 +1,122 @@ +Require Import Recdef. +Require Import FSetInterface. +Require Import InterfGraphMapImp. +Require Import OrderedOption. +Require Import FMapAVL. +Require Import IRC_termination. +Require Import IRC_graph. +Require Import IRC_Graph_Functions. +Require Import Edges. + +Import Edge RegFacts Props OTFacts. + +(* Definition of Register options *) + +Module OptionReg := OrderedOpt Register. + +(* Definition of the type of colorings *) + +Definition Coloring := Register.t -> OptionReg.t. + +(* Map used to build a coloring of the graph. + A coloring is a the function find applied to the map *) + +Module ColorMap := FMapAVL.Make Register. + +(* Function to complete a coloring by giving the + same color to coalesced vertices *) + +Definition complete_coloring e col : ColorMap.t Register.t := +let x := snd_ext e in let y := fst_ext e in +match ColorMap.find y col with +| None => col +| Some c => ColorMap.add x c col +end. + +(* Function to compute forbidden colors when completing the coloring + for simplified or spilled registers (optimistic spilling) *) + +Section add_monad. + +Definition vertex_add_monad (o : option Register.t) (VS : VertexSet.t) := +match o with +|Some v => VertexSet.add v VS +|None => VS +end. + +Variable A : Type. + +Lemma monad_fold : forall f a l s, +VertexSet.Equal (fold_left + (fun (x : VertexSet.t) (e : A) => vertex_add_monad (f e) x) + (a :: l) s) + (vertex_add_monad (f a) + (fold_left (fun (x : VertexSet.t) (e : A) => vertex_add_monad (f e) x) l s)). + +Proof. +intros f a l s;apply fold_left_assoc. +intros y z e;unfold vertex_add_monad. +destruct (f y);destruct (f z);[apply add_add|intuition|intuition|intuition]. +intros e1 e2 y H. unfold vertex_add_monad. +destruct (f y);intuition. +apply Dec.F.add_m; intuition. +Qed. + +End add_monad. + +Definition forbidden_colors x col g := +VertexSet.fold (fun v => vertex_add_monad (ColorMap.find v col)) + (interference_adj x g) + VertexSet.empty. + +(* Function to complete the coloring for simplified or spilled vertices. + Calls forbidden_colors. Is choosing yet any available color *) + +Definition available_coloring ircg x (col : ColorMap.t Register.t) := +let g := (irc_g ircg) in let palette := (pal ircg) in +match (VertexSet.choose (VertexSet.diff palette (forbidden_colors x col g))) with +| None => col +| Some c => ColorMap.add x c col +end. + +(* Definition of the empty coloring as the coloring where the only + colored vertices are the precolored ones *) + +Definition precoloring_map g := +VertexSet.fold (fun x => ColorMap.add x x) (precolored g) (ColorMap.empty Register.t). + +Function IRC_map (g : irc_graph) +{measure irc_measure g} : ColorMap.t Register.t := +match simplify g with +|Some rg' => let (r,g') := rg' in available_coloring g r (IRC_map g') +|None => match coalesce g with + |Some eg' => let (e,g') := eg' in complete_coloring e (IRC_map g') + |None => match freeze g with + |Some rg' => let (r,g') := rg' in IRC_map g' + |None => match spill g with + |Some rg' => let (r,g') := rg' in available_coloring g r (IRC_map g') + |None => precoloring_map (irc_g g) + end + end + end +end. + +Proof. +intros. apply (simplify_dec g r g' teq). +intros. apply (coalesce_dec g e g' teq0). +intros. apply (freeze_dec g r g' teq1). +intros. apply (spill_dec g r g' teq2). +Qed. + +(* Definition of the transformation from a map to a coloring, + simply by using find *) + +Definition map_to_coloring (colmap : ColorMap.t Register.t) := +fun x => ColorMap.find x colmap. + +(* Final definition of iterated register coalescing *) + +Definition IRC g palette := +let g' := graph_to_IRC_graph g palette in map_to_coloring (IRC_map g'). + +Definition precoloring g := map_to_coloring (precoloring_map g). diff --git a/backend/IRCColoring.v b/backend/IRCColoring.v new file mode 100755 index 00000000..5efeed7b --- /dev/null +++ b/backend/IRCColoring.v @@ -0,0 +1,847 @@ +Require Import Recdef. +Require Import IRC. +Require Import FSetInterface. +Require Import Edges. +Require Import Interference_adjacency. +Require Import IRC_graph. +Require Import Conservative_criteria. +Require Import InterfGraphMapImp. +Require Import Graph_Facts. +Require Import IRC_Graph_Functions. +Require Import WS. + +Import Edge RegFacts Props OTFacts. + +Module MapFacts := FMapFacts.Facts ColorMap. + +Definition proper_coloring_1 (col : Coloring) g := forall e x y, +interf_edge e -> +In_graph_edge e g -> +OptionReg.eq (col (fst_ext e)) (Some x) -> +OptionReg.eq (col (snd_ext e)) (Some y) -> +~Register.eq x y. + +Definition proper_coloring_2 (col : Coloring) g := forall x y, +OptionReg.eq (col x) (Some y) -> In_graph x g. + +Definition proper_coloring_3 (col : Coloring) palette := forall x y, +OptionReg.eq (col x) (Some y) -> VertexSet.In y palette. + +Definition proper_coloring col g palette := +proper_coloring_1 col g /\ proper_coloring_2 col g /\ proper_coloring_3 col palette. + +Lemma diff_empty_sub : forall s1 s2, +VertexSet.Empty (VertexSet.diff s1 s2) -> +VertexSet.Subset s1 s2. + +Proof. +intros s1 s2 H. +intros a H0. +destruct (In_dec a s2). +assumption. +elim (H a). +apply (VertexSet.diff_3 H0 n). +Qed. + +Lemma interf_in_forbidden : forall x y col g c, +Interfere x y g -> +OptionReg.eq (map_to_coloring col y) (Some c) -> +VertexSet.In c (forbidden_colors x col g). + +Proof. +intros x y col g c H H0. +unfold forbidden_colors. +unfold Interfere in H. rewrite <-in_interf in H. +generalize (interf_adj_comm _ _ _ H). intro H1. +generalize (VertexSet.elements_1 H1);clear H1;intro H1. +rewrite VertexSet.fold_1. +induction (VertexSet.elements (interference_adj x g)). +inversion H1. +rewrite monad_fold. +inversion H1;subst. +unfold map_to_coloring in H0. +inversion H0;subst. +rewrite (MapFacts.find_o col H3) in H2. +rewrite <-H2. +simpl;apply VertexSet.add_1. +assumption. +destruct (ColorMap.find a col);[apply VertexSet.add_2|];apply IHl;assumption. +Qed. + +Lemma available_coloring_1 : forall col x g, +proper_coloring (map_to_coloring col) (irc_g g) (pal g) -> +In_graph x (irc_g g) -> +proper_coloring (map_to_coloring (available_coloring g x col)) (irc_g g) (pal g). + +Proof. +intros col x ircg H HH. unfold available_coloring. +set (palette := pal ircg) in *. set (g := irc_g ircg) in *. +case_eq (VertexSet.choose (VertexSet.diff palette (forbidden_colors x col g))). +intros c H1. +generalize H1;intro H0. +unfold proper_coloring. +split. +unfold proper_coloring_1. +intros e x' y' H2 H3 H4 H5. +destruct (Register.eq_dec x (fst_ext e)). +unfold map_to_coloring in H4;unfold map_to_coloring in H5. +rewrite MapFacts.add_eq_o in H4. +rewrite MapFacts.add_neq_o in H5. +generalize (VertexSet.choose_1 H1);clear H1;intro H1. +generalize (VertexSet.diff_1 H1);intro H6. +generalize (VertexSet.diff_2 H1);intro H7;clear H1. +assert (Interfere x (snd_ext e) g) as Hinterf. +unfold Interfere. +assert (eq (fst_ext e, snd_ext e,None) (x, snd_ext e, None)). +Eq_eq. +rewrite (edge_eq e) in H3. +rewrite H2 in H3. +rewrite H1 in H3;assumption. +generalize (interf_in_forbidden x (snd_ext e) col g y' Hinterf H5);intro H1. +intro H8. +inversion H4;subst. +generalize (Register.eq_trans H11 H8);clear H11 H8;intro H8. +rewrite H8 in H7. +elim (H7 H1). +intro Heq. +elim (In_graph_edge_diff_ext _ _ H3). +apply (Register.eq_trans (Register.eq_sym Heq) e0). +assumption. +destruct (Register.eq_dec x (snd_ext e)). +unfold map_to_coloring in H4;unfold map_to_coloring in H5. +rewrite MapFacts.add_eq_o in H5. +rewrite MapFacts.add_neq_o in H4. +generalize (VertexSet.choose_1 H1);clear H1;intro H1. +generalize (VertexSet.diff_1 H1);intro H6. +generalize (VertexSet.diff_2 H1);intro H7;clear H1. +assert (Interfere x (fst_ext e) g) as Hinterf. +unfold Interfere. +rewrite (edge_eq e) in H3. rewrite H2 in H3. +assert (eq (fst_ext e, snd_ext e,None) (x, fst_ext e, None)) by Eq_comm_eq. +rewrite H1 in H3;assumption. +generalize (interf_in_forbidden x (fst_ext e) col g x' Hinterf H4);intro H1. +intro H8. +inversion H5;subst. +generalize (Register.eq_trans H11 (Register.eq_sym H8));clear H11 H8;intro H8. +rewrite H8 in H7. +elim (H7 H1). +intro Heq. +elim (In_graph_edge_diff_ext _ _ H3). +apply (Register.eq_trans (Register.eq_sym e0) Heq). +assumption. +unfold map_to_coloring in H4;unfold map_to_coloring in H5. +rewrite MapFacts.add_neq_o in H5. +rewrite MapFacts.add_neq_o in H4. +unfold proper_coloring in H;destruct H as [H _]. +unfold proper_coloring_1 in H. +apply (H e);assumption. +assumption. +assumption. +split. +unfold proper_coloring in *. +destruct H as [_ H];destruct H as [H _]. +unfold proper_coloring_2 in *. +intros x' y' H2. +destruct (Register.eq_dec x x'). +unfold map_to_coloring in H2. +rewrite e in HH;assumption. +apply (H x' y'). +unfold map_to_coloring in H2. +rewrite MapFacts.add_neq_o in H2. +assumption. +assumption. +unfold proper_coloring in *. +do 2 destruct H as [_ H]. +unfold proper_coloring_3 in *. +intros x' y' H2. +destruct (Register.eq_dec x x'). +unfold map_to_coloring in H2. +generalize (VertexSet.choose_1 H1);clear H1;intro H1. +generalize (VertexSet.diff_1 H1);intro H3. +rewrite MapFacts.add_eq_o in H2. +inversion H2;subst. +rewrite H6 in H3;assumption. +assumption. +apply (H x' y'). +unfold map_to_coloring in H2. +rewrite MapFacts.add_neq_o in H2. +assumption. +assumption. +auto. +Qed. + +Lemma complete_coloring_1 : forall col e g, +In_graph_edge e g -> +ColorMap.find (snd_ext e) col = None -> +OptionReg.eq +(map_to_coloring (complete_coloring e col) (snd_ext e)) +(map_to_coloring (complete_coloring e col) (fst_ext e)). + +Proof. +intros col e g HH Hin;unfold complete_coloring. +case_eq (ColorMap.find (fst_ext e) col). +intros r H. +unfold map_to_coloring. +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite H;apply OptionReg.eq_refl. +apply (In_graph_edge_diff_ext _ _ HH). +apply Register.eq_refl. +intro H0. +unfold map_to_coloring. +rewrite Hin;rewrite H0;apply OptionReg.eq_refl. +Qed. + +Lemma complete_coloring_2 : forall col e x, +~Register.eq (snd_ext e) x -> +OptionReg.eq +(map_to_coloring (complete_coloring e col) x) +(map_to_coloring col x). + +Proof. +intros col e x H. +unfold complete_coloring. +destruct (ColorMap.find (fst_ext e) col). +unfold map_to_coloring. +rewrite MapFacts.add_neq_o. +apply OptionReg.eq_refl. +assumption. +apply OptionReg.eq_refl. +Qed. + +Definition compat_col col := forall x y, +Register.eq x y -> OptionReg.eq (col x) (col y). + +Lemma compat_IRC : forall g palette, compat_col (IRC g palette). + +Proof. +unfold IRC. +unfold compat_col. +intros g palette x y H. +unfold map_to_coloring. +rewrite MapFacts.find_o with (y := y). +apply OptionReg.eq_refl. +assumption. +Qed. + +Lemma compat_complete : forall col e, +compat_col (map_to_coloring col) -> +compat_col (map_to_coloring (complete_coloring e col)). + +Proof. +unfold compat_col;intros col e H. +intros x y Heq;generalize (H x y Heq);clear H;intro H. +unfold complete_coloring. +destruct (ColorMap.find (fst_ext e) col). +unfold map_to_coloring. +destruct (Register.eq_dec x (snd_ext e)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +apply OptionReg.eq_refl. +apply (Register.eq_trans (Register.eq_sym e0) Heq). intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +assumption. +intro Helim;elim n. +apply (Register.eq_trans Heq (Register.eq_sym Helim)). +intro H0;elim n;auto. +assumption. +Qed. + +Lemma colored_complete_diff_colored : forall e col, +~Register.eq (snd_ext e) (fst_ext e) -> +OptionReg.eq (map_to_coloring (complete_coloring e col) (fst_ext e)) (map_to_coloring col (fst_ext e)). + +Proof. +intros e col H. +unfold complete_coloring. +case_eq (ColorMap.find (fst_ext e) col). +intros r H0. +unfold map_to_coloring. +rewrite MapFacts.add_neq_o. +apply OptionReg.eq_refl. +assumption. +intro H0. +apply OptionReg.eq_refl. +Qed. + +Lemma complete_coloring_coloring : forall col e g Hin Haff, +~VertexSet.In (snd_ext e) (precolored (irc_g g)) -> +proper_coloring (map_to_coloring col) (merge e (irc_g g) Hin Haff) (pal g) -> +conservative_coalescing_criteria e (irc_g g) (irc_k g) = true -> +compat_col (map_to_coloring col) -> +proper_coloring (map_to_coloring (complete_coloring e col)) (irc_g g) (pal g). + +Proof. +unfold proper_coloring;unfold proper_coloring_1; +unfold proper_coloring_2;unfold proper_coloring_3. +intros col e ircg H1 Haff Hpre H H0 Hcompat. +set (g := irc_g ircg) in *. set (palette := pal ircg) in *. +rewrite <-(Hk ircg) in *. +destruct H as [H HH];destruct HH as [HH HHH]. +split. +intros e' x' y' H2 H3 H4 H5. +apply (H (redirect (snd_ext e) (fst_ext e) e')). +unfold interf_edge;rewrite redirect_weight_eq;assumption. +apply In_merge_interf_edge. +assumption. +assumption. +(*intro Heq. +generalize (get_weight_m _ _ Heq);intro H6. +rewrite H2 in H6. +unfold aff_edge in Haff. +destruct Haff as [w Haff]. +rewrite Haff in H6. +inversion H6. +assumption.*) +destruct e' as [e' we'];destruct e' as [e'1 e'2]. +unfold redirect. +unfold interf_edge in H2;simpl in H2;subst. +change_rewrite. +destruct (OTFacts.eq_dec e'1 (snd_ext e));change_rewrite. +generalize (compat_complete col e);intro Hcompat2. +generalize (Hcompat2 Hcompat _ _ r);intro H7. +generalize (OptionReg.eq_trans _ _ _ (OptionReg.eq_sym _ _ H4) H7); +clear H7;intro H7. +assert (ColorMap.find (snd_ext e) col = None) as Hn. +case_eq (ColorMap.find (snd_ext e) col). +clear H3 H4 H5 r Hcompat2 H7. +intros r H8. +generalize (HH (snd_ext e) r);intro Helim. +unfold map_to_coloring in Helim. +rewrite H8 in Helim. +generalize (Helim (OptionReg.eq_refl _));clear Helim;intro Helim. +rewrite In_merge_vertex in Helim. destruct Helim. elim H3. auto. +auto. +generalize (complete_coloring_1 col e g H1 Hn);intro H8. +generalize (OptionReg.eq_trans _ _ _ H7 H8);intro H9. +generalize (complete_coloring_2 col e (fst_ext e) (In_graph_edge_diff_ext _ _ H1)). +intro H10. +generalize (OptionReg.eq_trans _ _ _ H9 H10);intro H11. +apply OptionReg.eq_sym;assumption. +destruct (OTFacts.eq_dec e'2 (snd_ext e));change_rewrite; +apply OptionReg.eq_trans with (y := (map_to_coloring (complete_coloring e col) e'1)); +[apply OptionReg.eq_sym;apply complete_coloring_2;intuition|assumption| +apply OptionReg.eq_sym;apply complete_coloring_2;intuition|assumption]. +destruct e' as [e' we'];destruct e' as [e'1 e'2]. +unfold redirect;change_rewrite. +unfold interf_edge in H2;simpl in H2;subst. +destruct (OTFacts.eq_dec e'1 (snd_ext e));change_rewrite. +destruct (Register.eq_dec (snd_ext e) e'2). +elim (In_graph_edge_diff_ext _ _ H3). +apply (Register.eq_trans (Register.eq_sym e0) (Register.eq_sym r)). +generalize (complete_coloring_2 col e e'2 n);intro H6. +apply OptionReg.eq_trans with (y := map_to_coloring (complete_coloring e col) e'2). +apply OptionReg.eq_sym;assumption. +assumption. +destruct (OTFacts.eq_dec e'2 (snd_ext e));change_rewrite. +generalize (compat_complete col e);intro Hcompat2. +generalize (Hcompat2 Hcompat _ _ r);intro H7. +generalize (OptionReg.eq_trans _ _ _ (OptionReg.eq_sym _ _ H5) H7); +clear H7;intro H7. +assert (ColorMap.find (snd_ext e) col = None) as Hn. +case_eq (ColorMap.find (snd_ext e) col). +clear H3 H4 H5 r Hcompat2 H7. +intros r H8. +generalize (HH (snd_ext e) r);intro Helim. +unfold map_to_coloring in Helim. +rewrite H8 in Helim. +generalize (Helim (OptionReg.eq_refl _));clear Helim;intro Helim. +rewrite In_merge_vertex in Helim. destruct Helim. elim H3. auto. +auto. +generalize (complete_coloring_1 col e g H1 Hn);intro H8. +generalize (OptionReg.eq_trans _ _ _ H7 H8);intro H9. +generalize (complete_coloring_2 col e (fst_ext e) (In_graph_edge_diff_ext _ _ H1)). +intro H10. +generalize (OptionReg.eq_trans _ _ _ H9 H10);intro H11. +apply OptionReg.eq_sym;assumption. +change (snd_ext (e'1,e'2,None)) with e'2. +assert (~Register.eq (snd_ext e) e'2) by intuition. +generalize (complete_coloring_2 col e e'2 H2);intro H6. +apply OptionReg.eq_trans with (y := map_to_coloring (complete_coloring e col) e'2). +apply OptionReg.eq_sym;assumption. +assumption. +(* second step *) +split. +intros x y H2. +destruct (Register.eq_dec x (snd_ext e)). +rewrite e0. +apply (In_graph_edge_in_ext _ _ H1). +assert (In_graph x (merge e g H1 Haff)). +apply HH with (y := y). +apply OptionReg.eq_trans with (y := map_to_coloring (complete_coloring e col) x). +apply OptionReg.eq_sym. +apply complete_coloring_2. +intuition. +assumption. +rewrite In_merge_vertex in H3. intuition. +(* third step *) +intros x y H2. +destruct (Register.eq_dec x (snd_ext e)). +assert (ColorMap.find (snd_ext e) col = None) as Hn. +case_eq (ColorMap.find (snd_ext e) col). +intros r H8. +generalize (HH (snd_ext e) r);intro Helim. +unfold map_to_coloring in Helim. +rewrite H8 in Helim. +generalize (Helim (OptionReg.eq_refl _));clear Helim;intro Helim. +rewrite In_merge_vertex in Helim. destruct Helim. elim H4. auto. +auto. +generalize (complete_coloring_1 col e g H1 Hn);intro H3. +generalize (complete_coloring_2 col e (fst_ext e) (In_graph_edge_diff_ext _ _ H1)). +intro H4. +generalize (OptionReg.eq_trans _ _ _ (OptionReg.eq_sym _ _ H4) (OptionReg.eq_sym _ _ H3)). +clear H3 H4;intro H3. +generalize (compat_complete col e Hcompat);intro Hcompat2. +generalize (Hcompat2 _ _ e0). +intro H4. +generalize (OptionReg.eq_trans _ _ _ H3 (OptionReg.eq_sym _ _ H4)). +clear H3 H4;intro H3. +apply HHH with (x:= fst_ext e). +apply (OptionReg.eq_trans _ _ _ H3 H2). +assert (~Register.eq (snd_ext e) x) by intuition. +generalize (complete_coloring_2 col e x H3). +intro H4. +apply HHH with (x:= x). +apply (OptionReg.eq_trans _ _ _ (OptionReg.eq_sym _ _ H4) H2). +Qed. + +Lemma interf_edge_in_delete_preference : forall x e g H, +interf_edge e -> +In_graph_edge e g -> +In_graph_edge e (delete_preference_edges x g H). + +Proof. +intros x e g HH H H0. +generalize (proj2 (In_graph_interf_edge_in_IE _ _) (conj H H0));intro H1. +rewrite In_delete_preference_edges_edge. +split. assumption. +intro. destruct H2. destruct H2. congruence. +Qed. + +Lemma coloring_delete_preference : forall col x g H palette, +proper_coloring col (delete_preference_edges x g H) palette -> +proper_coloring col g palette. + +Proof. +unfold proper_coloring;unfold proper_coloring_1; +unfold proper_coloring_2;unfold proper_coloring_3. +intros col x g Hdp palette H. +destruct H as [H HH];destruct HH as [HH HHH]. +split. +intros e x' y' H0 H1 H2 H3. +apply H with (e:=e). +assumption. +apply interf_edge_in_delete_preference;assumption. +assumption. +assumption. +split. +intros x' y' H0. rewrite <-(In_delete_preference_edges_vertex x' x g Hdp). +apply (HH x' y');auto. +assumption. +Qed. + +Lemma proper_remove_proper : forall col g x palette, +proper_coloring col (remove_vertex x g) palette -> +compat_col col -> +proper_coloring col g palette. + +Proof. +intros col g x palette H Hcompat. +unfold proper_coloring in *;unfold proper_coloring_1 in *; +unfold proper_coloring_2 in *;unfold proper_coloring_3 in *. +destruct H as [H HH];destruct HH as [HH HHH]. +split. +intros e x' y' H0 H1 H2 H3. +destruct (incident_dec e x). +destruct H4. +generalize (Hcompat _ _ H4);intro H5. +generalize (OptionReg.eq_trans _ _ _ H5 H2);clear H5;intro H5. +generalize (HH _ _ H5). intro. +rewrite In_remove_vertex in H6. destruct H6. elim H7. auto. +generalize (Hcompat _ _ H4);intro H5. +generalize (OptionReg.eq_trans _ _ _ H5 H3);clear H5;intro H5. +generalize (HH _ _ H5). intro. +rewrite In_remove_vertex in H6. destruct H6. elim H7. auto. +apply H with (e:=e); auto. +rewrite In_remove_edge; auto. +split. +intros x' y' H0. +generalize (HH x' y' H0);intro H1. +rewrite In_remove_vertex in H1. intuition. +assumption. +Qed. + +Section Fold_Facts. + +Variable A : Type. + +Lemma fold_left_compat_map : forall (f : ColorMap.t Register.t -> A -> ColorMap.t Register.t) l e e', +ColorMap.Equal e e' -> +(forall e1 e2 a, ColorMap.Equal e1 e2 -> ColorMap.Equal (f e1 a) (f e2 a)) -> +ColorMap.Equal (fold_left f l e) (fold_left f l e'). + +Proof. +intros f l. +induction l;simpl. +auto. +intros e e' H H0 H1. +apply (IHl (f e a) (f e' a)). +apply H0;assumption. +assumption. +Qed. + +End Fold_Facts. + +Lemma empty_coloring_simpl : forall l a, +NoDupA Register.eq (a :: l) -> +ColorMap.Equal (fold_left (fun (a0 : ColorMap.t VertexSet.elt) (e : VertexSet.elt) => + ColorMap.add e e a0) (a :: l) (ColorMap.empty Register.t)) + (ColorMap.add a a ( + fold_left (fun (a0 : ColorMap.t VertexSet.elt) (e : VertexSet.elt) => + ColorMap.add e e a0) l (ColorMap.empty Register.t))). + +Proof. +intro l;generalize (ColorMap.empty Register.t). +induction l;simpl;intros. +unfold ColorMap.Equal;auto. +assert (ColorMap.Equal (ColorMap.add a a (ColorMap.add a0 a0 t0)) + (ColorMap.add a0 a0 (ColorMap.add a a t0))). +unfold ColorMap.Equal. +intro y. +destruct (Register.eq_dec a0 a). +inversion H;subst. +elim H2. +left;auto. +destruct (Register.eq_dec y a). +destruct (Register.eq_dec y a0). +elim (n (Register.eq_trans (Register.eq_sym e0) e)). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +reflexivity. +intuition. +intro Hneq;elim n0;auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Register.eq_dec a0 y). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +reflexivity. +assumption. +assumption. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +reflexivity. +intro Hneq;elim n0;auto. +assumption. +assumption. +intro Hneq;elim n0;auto. +rewrite (fold_left_compat_map _ + (fun (a1 : ColorMap.t VertexSet.elt) (e : VertexSet.elt) => ColorMap.add e e a1) l _ _ H0). +simpl in IHl;apply IHl. +constructor. +inversion H;subst. +intro H5;elim H3. +right;auto. +inversion H;subst. +inversion H4;subst;assumption. +intros e1 e2 a1 H1. +unfold ColorMap.Equal;intro y. +destruct (Register.eq_dec a1 y). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +reflexivity. +assumption. +assumption. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H1. +assumption. +assumption. +Qed. + +Lemma map_to_coloring_ext : forall c1 c2 g palette, +ColorMap.Equal c2 c1 -> +proper_coloring (map_to_coloring c1) g palette -> +proper_coloring (map_to_coloring c2) g palette. + +Proof. +unfold proper_coloring;unfold proper_coloring_1; +unfold proper_coloring_2;unfold proper_coloring_3. +intros c1 c2 g palette H H0. +destruct H0 as [H0 H1];destruct H1 as [H1 H2]. +split. +intros e x y H3 H4 H5 H6. +apply (H0 e);try assumption. +unfold ColorMap.Equal in H. +apply OptionReg.eq_trans with (y := (map_to_coloring c2 (fst_ext e))). +unfold map_to_coloring. +rewrite H. +apply OptionReg.eq_refl. +assumption. +apply OptionReg.eq_trans with (y := (map_to_coloring c2 (snd_ext e))). +unfold map_to_coloring. +rewrite H. +apply OptionReg.eq_refl. +assumption. +split. +intros x y H3. +apply (H1 x y). +apply OptionReg.eq_trans with (y := (map_to_coloring c2 x)). +unfold map_to_coloring. +rewrite H. +apply OptionReg.eq_refl. +assumption. +intros x y H3. +apply (H2 x y). +apply OptionReg.eq_trans with (y := (map_to_coloring c2 x)). +unfold map_to_coloring. +rewrite H. +apply OptionReg.eq_refl. +assumption. +Qed. + +Lemma compat_map_to_coloring : forall c1 c2 x, +ColorMap.Equal c1 c2 -> +OptionReg.eq (map_to_coloring c1 x) (map_to_coloring c2 x). + +Proof. +intros c1 c2 x H;unfold map_to_coloring; +unfold ColorMap.Equal in H. +rewrite H;apply OptionReg.eq_refl. +Qed. + +Lemma proper_coloring_empty_aux : forall g x, +OptionReg.eq (precoloring g x) (Some x) \/ +OptionReg.eq (precoloring g x) None. + +Proof. +intros g x. +unfold precoloring, precoloring_map. +rewrite VertexSet.fold_1. + +assert (NoDupA Register.eq (VertexSet.elements (precolored g))) as HNoDup by + (apply RegFacts.NoDupA_elements). +induction (VertexSet.elements (precolored g)). +simpl. +right;unfold map_to_coloring;rewrite MapFacts.empty_o. +apply OptionReg.eq_refl. +(* induction case *) +generalize (empty_coloring_simpl l a HNoDup);intro H0. +inversion HNoDup;subst. +generalize (IHl H3);clear IHl H2 H3;intro IHl. +destruct IHl. +left. +generalize (compat_map_to_coloring _ _ x H0);clear H0;intro H0. +apply (OptionReg.eq_trans _ _ _ H0). +destruct (Register.eq_dec a x). +unfold map_to_coloring. +rewrite MapFacts.add_eq_o. +constructor;assumption. +assumption. +unfold map_to_coloring. +rewrite MapFacts.add_neq_o. +apply H. +assumption. +destruct (Register.eq_dec a x). +left. +generalize (compat_map_to_coloring _ _ x H0);clear H0;intro H0. +apply (OptionReg.eq_trans _ _ _ H0). +unfold map_to_coloring. +rewrite MapFacts.add_eq_o. +constructor;assumption. +assumption. +right. +generalize (compat_map_to_coloring _ _ x H0);clear H0;intro H0. +apply (OptionReg.eq_trans _ _ _ H0). +unfold map_to_coloring. +rewrite MapFacts.add_neq_o. +apply H. +assumption. +Qed. + +Lemma in_empty_coloring_in_precolored : forall x g y, +OptionReg.eq (precoloring g x)(Some y) -> +VertexSet.In x (precolored g). + +Proof. +intros x g y H. unfold precoloring, precoloring_map in H. +rewrite VertexSet.fold_1 in H. +generalize (VertexSet.elements_2);intro H0. +generalize (H0 (precolored g));clear H0;intro H0. +assert (NoDupA Register.eq (VertexSet.elements (precolored g))) as HNoDup by + (apply RegFacts.NoDupA_elements). +induction (VertexSet.elements (precolored g)). +simpl in H. +unfold map_to_coloring in H;rewrite MapFacts.empty_o in H. +inversion H. +destruct (Register.eq_dec a x). +apply H0. +left; intuition. +apply IHl. +generalize (empty_coloring_simpl l a HNoDup);intro H1. +generalize (compat_map_to_coloring _ _ x H1);clear H1;intro H1. +unfold map_to_coloring in H1. +rewrite MapFacts.add_neq_o in H1. +apply (OptionReg.eq_trans _ _ _ (OptionReg.eq_sym _ _ H1)). +clear H1 IHl. +assumption. +assumption. +intros;apply H0;right;auto. +inversion HNoDup;auto. +Qed. + +Lemma proper_coloring_pre_inv : forall g x, +VertexSet.In x (precolored g) -> +OptionReg.eq (precoloring g x) (Some x). + +Proof. +intros g x H. unfold precoloring, precoloring_map. +rewrite VertexSet.fold_1. +generalize (VertexSet.elements_1);intro H0. +generalize (H0 (precolored g) x);clear H0;intro H0. +assert (NoDupA Register.eq (VertexSet.elements (precolored g))) as HNoDup by + (apply RegFacts.NoDupA_elements). +induction (VertexSet.elements (precolored g)). +generalize (H0 H). intro. inversion H1. +generalize (empty_coloring_simpl l a HNoDup);intro H1. simpl in *. +generalize (compat_map_to_coloring _ _ x H1). intro. +unfold map_to_coloring in H2. +destruct (Register.eq_dec x a). + rewrite MapFacts.add_eq_o in H2. +apply (OptionReg.eq_trans _ _ _ H2). +constructor. apply Register.eq_sym. auto. +apply Register.eq_sym. auto. +rewrite (MapFacts.add_neq_o) in H2. apply (OptionReg.eq_trans _ _ _ H2). +apply IHl. clear H1 H2. +intros. generalize (H0 H1). intro. +inversion H2; subst. elim n. auto. +auto. +inversion HNoDup;auto. +auto. +Qed. + +Lemma proper_coloring_empty : forall g palette, +VertexSet.Subset (precolored g) palette -> +proper_coloring (precoloring g) g palette. + +Proof. +intros g palette H. +unfold proper_coloring;unfold proper_coloring_1; +unfold proper_coloring_2;unfold proper_coloring_3. +split;[|split]. +intros e x y H0 H1 H2 H3 H4. +destruct (proper_coloring_empty_aux g (fst_ext e)); +generalize (OptionReg.eq_trans _ _ _ (OptionReg.eq_sym _ _ H2) H5); +clear H2 H5;intro H2. +destruct (proper_coloring_empty_aux g (snd_ext e)); +generalize (OptionReg.eq_trans _ _ _ (OptionReg.eq_sym _ _ H3) H5); +clear H3 H5;intro H3. +elim (In_graph_edge_diff_ext _ _ H1). +inversion H2;inversion H3;subst. +apply (Register.eq_trans (Register.eq_sym H10) (Register.eq_trans (Register.eq_sym H4) H7)). +inversion H3. +inversion H2. +intros x y H0. +assert (VertexSet.In x (precolored g)). +apply in_empty_coloring_in_precolored with (y := y). assumption. +rewrite precolored_equiv in H1. intuition. +intros x y H0. +apply H. +destruct (proper_coloring_empty_aux g x); +generalize (OptionReg.eq_trans _ _ _ (OptionReg.eq_sym _ _ H1) H0); +intro H2;inversion H2;subst. +rewrite <-H5. +apply in_empty_coloring_in_precolored with (y := x). +assumption. +Qed. + +Lemma proper_coloring_IRC_map : forall gp, +VertexSet.Subset (precolored (irc_g gp)) (pal gp) -> +proper_coloring (map_to_coloring (IRC_map gp)) (irc_g gp) (pal gp). + +Proof. +intros gp Hpre. +functional induction IRC_map gp;simpl in *. + +(* simplify *) +generalize (simplify_inv _ _ e). intro. +generalize (simplify_inv2 _ _ e). intro. destruct H0. simpl in H0. +apply available_coloring_1. +apply proper_remove_proper with (x:=r). +rewrite H0 in *. unfold simplify_irc in *. simpl in *. +apply IHt0. rewrite precolored_remove_vertex. +unfold VertexSet.Subset. intros. apply Hpre. apply (VertexSet.remove_3 H1). + +unfold compat_col. +intros. unfold map_to_coloring. +rewrite (MapFacts.find_o _ H1). apply OptionReg.eq_refl. + +apply (In_simplify_props _ _ _ _ _ _ _ _ (VertexSet.choose_1 H) (refl_equal _) (HWS_irc g)). + +(*coalesce*) +generalize (coalesce_inv _ _ e0). intro. +generalize (coalesce_inv_2 _ _ e0). intro. +destruct H0. destruct H0. simpl in *. + +assert (forall e', EdgeSet.In e' (get_movesWL (irc_wl g)) -> In_graph_edge e' (irc_g g)). +intros. generalize (In_move_props _ _ _ _ _ _ _ _ H1 (refl_equal _) (HWS_irc g)). intuition. + +apply (complete_coloring_coloring _ _ _ x x0). +apply (any_coalescible_edge_2 e1 _ _ (get_movesWL (irc_wl g)) H1 H). +rewrite H0 in *. unfold merge_irc in *. simpl in *. apply IHt0. +rewrite precolored_merge. +unfold VertexSet.Subset. intros. apply Hpre. apply (VertexSet.remove_3 H2). +apply (proj1 (any_coalescible_edge_1 _ _ _ _ H1 H)). + +unfold compat_col. +intros. unfold map_to_coloring. +rewrite (MapFacts.find_o _ H2). apply OptionReg.eq_refl. + +(* freeze *) +generalize (freeze_inv _ _ e1). intro. +generalize (freeze_inv2 _ _ e1). intro. destruct H0. destruct H0. simpl in *. +apply (coloring_delete_preference _ r (irc_g g) x0 _). +rewrite H0 in *. unfold delete_preference_edges_irc2 in *. simpl in *. apply IHt0. +unfold VertexSet.Subset. intros. apply Hpre. +rewrite precolored_delete_preference_edges in H1. assumption. + +(* spill *) +generalize (spill_inv _ _ e2). intro. +generalize (spill_inv2 _ _ e2). intro. destruct H0. simpl in *. +apply available_coloring_1. +apply proper_remove_proper with (x:=r). +rewrite H0 in *. unfold spill_irc in *. simpl in *. apply IHt0. +unfold VertexSet.Subset. intros. apply Hpre. +rewrite precolored_remove_vertex in H1. apply (VertexSet.remove_3 H1). + +unfold compat_col. +intros. unfold map_to_coloring. +rewrite (MapFacts.find_o _ H1). apply OptionReg.eq_refl. + +apply (In_spill_props _ _ _ _ _ _ _ _ (lowest_cost_in _ _ _ H) (refl_equal _) (HWS_irc g)). + +(* terminal case *) +apply proper_coloring_empty. assumption. +Qed. + +Lemma proper_coloring_IRC_aux : forall g palette, +VertexSet.Subset (precolored g) palette -> +proper_coloring (IRC g palette) g palette. + +Proof. +intros. apply proper_coloring_IRC_map. auto. +Qed. + +Lemma proper_coloring_IRC : forall g palette, +proper_coloring (precoloring g) g palette -> +proper_coloring (IRC g palette) g palette. + +Proof. +intros. apply proper_coloring_IRC_map. +intro. intro. unfold proper_coloring, proper_coloring_3 in H. +do 2 destruct H as [_ H]. unfold graph_to_IRC_graph in *. simpl in *. +apply H with (x:= a). apply proper_coloring_pre_inv. assumption. +Qed. \ No newline at end of file diff --git a/backend/IRC_Graph_Functions.v b/backend/IRC_Graph_Functions.v new file mode 100755 index 00000000..fc691aa2 --- /dev/null +++ b/backend/IRC_Graph_Functions.v @@ -0,0 +1,548 @@ +Require Import FSets. +Require Import InterfGraphMapImp. +Require Import Spill_WL. +Require Import ZArith. +Require Import Simplify_WL. +Require Import Spill_WL. +Require Import Merge_WL. +Require Import Freeze_WL. +Require Import IRC_graph. +Require Import Edges. +Require Import Conservative_criteria. +Require Import WS. + +Import RegFacts Props OTFacts. + +Definition any_vertex := VertexSet.choose. + +(* simplify *) + +Definition simplify_irc r ircg H := +Make_IRC_Graph (remove_vertex r (irc_g ircg)) + (simplify_wl r ircg (irc_k ircg)) + (pal ircg) + (irc_k ircg) + (WS_simplify r ircg H) + (Hk ircg). + +Definition simplify g : option (Register.t * irc_graph) := +let simplifyWL := get_simplifyWL (irc_wl g) in +match any_vertex simplifyWL as v return (any_vertex simplifyWL = v -> option (Register.t * irc_graph)) with +|Some r => fun H : any_vertex simplifyWL = Some r => + Some (r, simplify_irc r g (VertexSet.choose_1 H)) +|None => fun H : any_vertex simplifyWL = None => None +end (refl_equal (any_vertex simplifyWL)). + +Lemma simplify_inv_aux : + forall g P, + match simplify g with + | Some x => + forall ( H : (any_vertex (get_simplifyWL (irc_wl g)) = Some (fst x))), + (simplify_irc (fst x) g (VertexSet.choose_1 H) = snd x) -> P + | None => + (any_vertex (get_simplifyWL (irc_wl g)) = None -> P) + end -> P. +Proof. + intros g P. + unfold simplify. + + set (simplifyWL := get_simplifyWL (irc_wl g)) in *. + set (Z := any_vertex simplifyWL) in *. + + refine + (match Z as W + return forall (H : Z = W), + +match + match W as v return (Z = v -> option (Register.t * irc_graph)) with + | Some r => + fun H : Z = Some r => Some (r, simplify_irc r g (VertexSet.choose_1 H)) + | None => fun _ : Z = None => None + end H +with +| Some x => forall H : Z = Some (fst x), simplify_irc (fst x) g (VertexSet.choose_1 H) = snd x -> P +| None => Z = None -> P +end -> P + + with + | Some x => _ + | None => _ + end _). + +simpl. intros. apply X with (H0 := H). reflexivity. +auto. +Qed. + +Lemma simplify_inv : forall g res, +simplify g = Some res -> +any_vertex (get_simplifyWL (irc_wl g)) = Some (fst res). +Proof. + intros. + apply simplify_inv_aux with g. + rewrite H. + auto. +Qed. + +Lemma simplify_inv2 : forall g res, +simplify g = Some res -> +exists H, snd res = simplify_irc (fst res) g (VertexSet.choose_1 H). + +Proof. +intros. +apply (simplify_inv_aux g). rewrite H. +simpl. intros. rewrite <-H1. +exists H0. reflexivity. +Qed. + +(* merge *) + +Definition merge_irc e ircg pin paff := +let g' := merge e (irc_g ircg) pin paff in +Make_IRC_Graph g' + (merge_wl3 e ircg g' pin paff) + (pal ircg) + (irc_k ircg) + (WS_coalesce _ _ pin paff) + (Hk ircg). + +Definition coalesce g : option (Edge.t * irc_graph) := +let movesWL := get_movesWL (irc_wl g) in +let graph := irc_g g in +let HWS := HWS_irc g in +let k := irc_k g in +match any_coalescible_edge movesWL graph k as e +return (any_coalescible_edge movesWL graph k = e -> option (Edge.t * irc_graph)) with +|Some edge => fun H : any_coalescible_edge movesWL graph k = Some edge => + let Hin := any_coalescible_edge_11 _ _ _ _ H in + let Hing := proj2 (In_move_props _ _ _ _ _ _ _ _ Hin (refl_equal _) HWS) in + let Haff := proj1 (In_move_props _ _ _ _ _ _ _ _ Hin (refl_equal _) HWS) in + Some (edge,merge_irc edge g Hing Haff) +|None => fun H : any_coalescible_edge movesWL graph k = None => None +end (refl_equal (any_coalescible_edge movesWL graph k)). + +Lemma coalesce_inv_aux : + forall g P, + match coalesce g with + | Some x => + forall (H : (any_coalescible_edge (get_movesWL (irc_wl g)) (irc_g g) (irc_k g) = Some (fst x))), + (merge_irc (fst x) g + (proj2 (In_move_props _ _ _ _ _ _ _ _ (any_coalescible_edge_11 _ _ _ _ H) (refl_equal _) (HWS_irc g))) + (proj1 (In_move_props _ _ _ _ _ _ _ _ (any_coalescible_edge_11 _ _ _ _ H) (refl_equal _) (HWS_irc g)))) + = snd x -> P + | None => + (any_coalescible_edge (get_movesWL (irc_wl g)) (irc_g g) (irc_k g) = None -> P) + end -> P. +Proof. + intros g P. + unfold coalesce. + + set (movesWL := get_movesWL (irc_wl g)) in *. + set (Z := any_coalescible_edge movesWL (irc_g g) (irc_k g)) in *. + + refine + (match Z as W + return forall (H : Z = W), + +match + match W as e return (Z = e -> option (Edge.t * irc_graph)) with + | Some edge => + fun H : Z = Some edge => Some (edge, +merge_irc edge g + (proj2 + (In_move_props edge (irc_g g) + (get_spillWL (irc_wl g), get_freezeWL (irc_wl g), + get_simplifyWL (irc_wl g), movesWL) movesWL + (get_spillWL (irc_wl g)) (get_freezeWL (irc_wl g)) + (get_simplifyWL (irc_wl g)) (irc_k g) + (any_coalescible_edge_11 edge (irc_g g) (irc_k g) movesWL H) + (refl_equal + (get_spillWL (irc_wl g), get_freezeWL (irc_wl g), + get_simplifyWL (irc_wl g), movesWL)) (HWS_irc g))) + (proj1 + (In_move_props edge (irc_g g) + (get_spillWL (irc_wl g), get_freezeWL (irc_wl g), + get_simplifyWL (irc_wl g), movesWL) movesWL + (get_spillWL (irc_wl g)) (get_freezeWL (irc_wl g)) + (get_simplifyWL (irc_wl g)) (irc_k g) + (any_coalescible_edge_11 edge (irc_g g) (irc_k g) movesWL H) + (refl_equal + (get_spillWL (irc_wl g), get_freezeWL (irc_wl g), + get_simplifyWL (irc_wl g), movesWL)) (HWS_irc g)))) + | None => fun _ : Z = None => None + end H +with +| Some x => + forall H : Z = Some (fst x), + merge_irc (fst x) g + (proj2 + (In_move_props (fst x) (irc_g g) + (get_spillWL (irc_wl g), get_freezeWL (irc_wl g), + get_simplifyWL (irc_wl g), movesWL) movesWL + (get_spillWL (irc_wl g)) (get_freezeWL (irc_wl g)) + (get_simplifyWL (irc_wl g)) (irc_k g) + (any_coalescible_edge_11 (fst x) (irc_g g) (irc_k g) movesWL H) + (refl_equal + (get_spillWL (irc_wl g), get_freezeWL (irc_wl g), + get_simplifyWL (irc_wl g), movesWL)) (HWS_irc g))) + (proj1 + (In_move_props (fst x) (irc_g g) + (get_spillWL (irc_wl g), get_freezeWL (irc_wl g), + get_simplifyWL (irc_wl g), movesWL) movesWL + (get_spillWL (irc_wl g)) (get_freezeWL (irc_wl g)) + (get_simplifyWL (irc_wl g)) (irc_k g) + (any_coalescible_edge_11 (fst x) (irc_g g) (irc_k g) movesWL H) + (refl_equal + (get_spillWL (irc_wl g), get_freezeWL (irc_wl g), + get_simplifyWL (irc_wl g), movesWL)) (HWS_irc g))) = snd x -> + P +| None => Z = None -> P +end -> P + with + | Some x => _ + | None => _ + end _). + +simpl. intros. apply X with (H0 := H). reflexivity. +auto. +Qed. + +Lemma coalesce_inv : forall g res, +coalesce g = Some res -> +any_coalescible_edge (get_movesWL (irc_wl g)) (irc_g g) (irc_k g) = Some (fst res). +Proof. + intros. + apply (coalesce_inv_aux g). + rewrite H. + auto. +Qed. + +Lemma coalesce_inv_2 : forall g res, +coalesce g = Some res -> +exists H, exists H', snd res = merge_irc (fst res) g H H'. + +Proof. +intros. +apply (coalesce_inv_aux g). +rewrite H. +simpl. intros. +exists ((proj2 + (In_move_props (fst res) (irc_g g) + (get_spillWL (irc_wl g), get_freezeWL (irc_wl g), + get_simplifyWL (irc_wl g), get_movesWL (irc_wl g)) + (get_movesWL (irc_wl g)) (get_spillWL (irc_wl g)) + (get_freezeWL (irc_wl g)) (get_simplifyWL (irc_wl g)) + (irc_k g) + (any_coalescible_edge_11 (fst res) (irc_g g) (irc_k g) + (get_movesWL (irc_wl g)) H0) + (refl_equal + (get_spillWL (irc_wl g), get_freezeWL (irc_wl g), + get_simplifyWL (irc_wl g), get_movesWL (irc_wl g))) + (HWS_irc g)))). +exists ((proj1 + (In_move_props (fst res) (irc_g g) + (get_spillWL (irc_wl g), get_freezeWL (irc_wl g), + get_simplifyWL (irc_wl g), get_movesWL (irc_wl g)) + (get_movesWL (irc_wl g)) (get_spillWL (irc_wl g)) + (get_freezeWL (irc_wl g)) (get_simplifyWL (irc_wl g)) + (irc_k g) + (any_coalescible_edge_11 (fst res) (irc_g g) (irc_k g) + (get_movesWL (irc_wl g)) H0) + (refl_equal + (get_spillWL (irc_wl g), get_freezeWL (irc_wl g), + get_simplifyWL (irc_wl g), get_movesWL (irc_wl g))) + (HWS_irc g)))). +auto. +Qed. + +(* freeze *) + +Definition delete_preference_edges_irc2 v ircg Hing Hdep := +let k := irc_k ircg in +let g' := delete_preference_edges v (irc_g ircg) Hing in +Make_IRC_Graph g' + (delete_preferences_wl2 v ircg k) + (pal ircg) + (irc_k ircg) + (WS_freeze v ircg Hing Hdep) + (Hk ircg). + +Definition freeze g : option (Register.t * irc_graph) := +let freezeWL := get_freezeWL (irc_wl g) in +let graph := irc_g g in +let HWS := HWS_irc g in +match any_vertex freezeWL as r +return (any_vertex freezeWL = r -> option (Register.t*irc_graph)) with +|Some x => fun H : any_vertex freezeWL = Some x => + let Hin := VertexSet.choose_1 H in + let Hing := proj1 (proj2 (proj2 (In_freeze_props _ _ _ _ _ _ _ _ Hin (refl_equal _) HWS)))in + Some (x,delete_preference_edges_irc2 x g Hing Hin) +|None => fun H : any_vertex freezeWL = None => None +end (refl_equal (any_vertex freezeWL)). + +Lemma freeze_inv_aux : + forall g P, + match freeze g with + | Some x => + forall ( H : (any_vertex (get_freezeWL (irc_wl g)) = Some (fst x))), + (delete_preference_edges_irc2 (fst x) g + (proj1 (proj2 (proj2 (In_freeze_props _ _ _ _ _ _ _ _ (VertexSet.choose_1 H) (refl_equal _) (HWS_irc g))))) + (VertexSet.choose_1 H) = snd x) -> P + | None => + (any_vertex (get_freezeWL (irc_wl g)) = None -> P) + end -> P. + +Proof. + intros g P. + unfold freeze. + + set (freezeWL := get_freezeWL (irc_wl g)) in *. + set (Z := any_vertex freezeWL) in *. + + refine + (match Z as W + return forall (H : Z = W), + +match + match W as v return (Z = v -> option (Register.t * irc_graph)) with + | Some x => + fun H : Z = Some x => Some + (x, + delete_preference_edges_irc2 x g + (proj1 + (proj2 + (proj2 + (In_freeze_props x (irc_g g) + (get_spillWL (irc_wl g), freezeWL, + get_simplifyWL (irc_wl g), get_movesWL (irc_wl g)) + freezeWL (get_spillWL (irc_wl g)) + (get_simplifyWL (irc_wl g)) (get_movesWL (irc_wl g)) + (irc_k g) (VertexSet.choose_1 H) + (refl_equal + (get_spillWL (irc_wl g), freezeWL, + get_simplifyWL (irc_wl g), get_movesWL (irc_wl g))) + (HWS_irc g))))) (VertexSet.choose_1 H)) + | None => fun _ : Z = None => None + end H +with +| Some x => + +forall H : Z = Some (fst x), + delete_preference_edges_irc2 (fst x) g + (proj1 + (proj2 + (proj2 + (In_freeze_props (fst x) (irc_g g) + (get_spillWL (irc_wl g), freezeWL, + get_simplifyWL (irc_wl g), get_movesWL (irc_wl g)) freezeWL + (get_spillWL (irc_wl g)) (get_simplifyWL (irc_wl g)) + (get_movesWL (irc_wl g)) (irc_k g) + (VertexSet.choose_1 H) + (refl_equal + (get_spillWL (irc_wl g), freezeWL, + get_simplifyWL (irc_wl g), get_movesWL (irc_wl g))) + (HWS_irc g))))) (VertexSet.choose_1 H) = snd x -> P +| None => Z = None -> P +end -> P + with + | Some x => _ + | None => _ + end _). + +simpl. intros. apply X with (H0 := H). reflexivity. +auto. +Qed. + +Lemma freeze_inv : forall g res, +freeze g = Some res -> +any_vertex (get_freezeWL (irc_wl g)) = Some (fst res). +Proof. + intros. + apply freeze_inv_aux with g. + rewrite H. + auto. +Qed. + +Lemma freeze_inv2 : forall g res, +freeze g = Some res -> +exists H', exists H, snd res = delete_preference_edges_irc2 (fst res) g H H'. + +Proof. +intros. +apply (freeze_inv_aux g). rewrite H. +simpl. intros. rewrite <-H1. +exists (VertexSet.choose_1 H0). +exists (proj1 (proj2 (proj2 (In_freeze_props _ _ _ _ _ _ _ _ (VertexSet.choose_1 H0) (refl_equal _) (HWS_irc g))))). reflexivity. +Qed. + +(* spill *) +Definition spill_irc r ircg H := +Make_IRC_Graph (remove_vertex r (irc_g ircg)) + (spill_wl r ircg (irc_k ircg)) + (pal ircg) + (irc_k ircg) + (WS_spill r ircg H) + (Hk ircg). + +Definition cost_order (opt : (Register.t*nat*nat)) y g := +let (tmp, pref_card) := opt in +let (x, int_card) := tmp in +let y_int := VertexSet.cardinal (interference_adj y g) in +match lt_eq_lt_dec y_int int_card with +|inleft (left _) => opt +|inleft (right _) => let y_pref := VertexSet.cardinal (preference_adj y g) in + match le_lt_dec pref_card y_pref with + |left _ => opt + |right _ => (y, y_int, y_pref) + end +|inright _ => (y, y_int, VertexSet.cardinal (preference_adj y g)) +end. + +Definition lowest_cost_aux s g o := +VertexSet.fold (fun v o => match o with + | Some opt => Some (cost_order opt v g) + | None => Some (v, VertexSet.cardinal (interference_adj v g), + VertexSet.cardinal (preference_adj v g)) + end) + s + o. + +Definition lowest_cost s g := +match lowest_cost_aux s g None with +| Some r => Some (fst (fst r)) +| None => None +end. + +Lemma lowest_cost_aux_in : forall x s g o, +lowest_cost_aux s g o = Some x-> +VertexSet.In (fst (fst x)) s \/ o = Some x. + +Proof. +intros. unfold lowest_cost_aux in H. +set (f := (fun (v : VertexSet.elt) (o : option (MyRegisters.Regs.t * nat * nat)) => + match o with + | Some opt => Some (cost_order opt v g) + | None => + Some + (v, VertexSet.cardinal (interference_adj v g), + VertexSet.cardinal (preference_adj v g)) + end )) in *. +unfold VertexSet.elt in *. +fold f in H. +rewrite VertexSet.fold_1 in H. +set (f' := fun a e => f e a) in *. +unfold VertexSet.elt in *. fold f' in H. +generalize VertexSet.elements_2. intro HH. +generalize (HH s). clear HH. intro HH. +generalize x o H. clear H. +induction (VertexSet.elements s). intros x0 o0 H. +simpl in H. right. auto. +simpl. intros. +assert (VertexSet.In (fst (fst x0)) s \/ (f' o0 a) = Some x0). +apply IHl. +intros. apply HH. right. auto. +auto. +destruct H0. +left. auto. +unfold f' in H0. unfold f in H0. +case_eq o0; intros. +rewrite H1 in *. +case_eq (cost_order p a g); intros. +rewrite H2 in H0. unfold cost_order in H2. +destruct p. simpl in *. destruct p. simpl in *. +destruct (lt_eq_lt_dec (VertexSet.cardinal (interference_adj a g)) n1). destruct s0. +right. rewrite H2. auto. +destruct (le_lt_dec n0 (VertexSet.cardinal (preference_adj a g))). +right. rewrite H2. auto. +left. apply HH. left. destruct p0. simpl in *. +destruct x0. destruct p. simpl in *. +inversion H0. inversion H2. subst. intuition. +left. apply HH. left. destruct p0. simpl in *. +destruct x0. destruct p. simpl in *. +inversion H0. inversion H2. subst. intuition. +rewrite H1 in H0. +unfold f' in H. rewrite H1 in H. simpl in H. rewrite H0 in H. +fold f' in H. +left. apply HH. inversion H0. simpl. left. intuition. +Qed. + +Lemma lowest_cost_in : forall x s g, +lowest_cost s g = Some x -> +VertexSet.In x s. + +Proof. +intros. unfold lowest_cost in H. +case_eq (lowest_cost_aux s g None); intros; rewrite H0 in H. +generalize (lowest_cost_aux_in p s g None H0). intro. +destruct H1. inversion H. subst. auto. +congruence. +congruence. +Qed. + +Definition spill g : option (Register.t * irc_graph) := +let spillWL := get_spillWL (irc_wl g) in +match lowest_cost spillWL (irc_g g) as v return (lowest_cost spillWL (irc_g g) = v -> option (Register.t * irc_graph)) with +|Some r => fun H : lowest_cost spillWL (irc_g g) = Some r => + Some (r, spill_irc r g (lowest_cost_in _ _ _ H)) +|None => fun H : lowest_cost spillWL (irc_g g) = None => None +end (refl_equal (lowest_cost spillWL (irc_g g))). + +Lemma spill_inv_aux : + forall g P, + match spill g with + | Some x => + forall ( H : (lowest_cost (get_spillWL (irc_wl g)) (irc_g g) = Some (fst x))), + (spill_irc (fst x) g (lowest_cost_in _ _ _ H) = snd x) -> P + | None => + (lowest_cost (get_spillWL (irc_wl g)) (irc_g g) = None -> P) + end -> P. +Proof. + intros g P. + unfold spill. + + set (spillWL := get_spillWL (irc_wl g)) in *. + set (Z := lowest_cost spillWL (irc_g g)) in *. + + refine + (match Z as W + return forall (H : Z = W), + +match + match W as v return (Z = v -> option (Register.t * irc_graph)) with + | Some r => + fun H : Z = Some r => Some (r, spill_irc r g (lowest_cost_in _ _ _ H)) + | None => fun _ : Z = None => None + end H +with +| Some x => forall H : Z = Some (fst x), spill_irc (fst x) g (lowest_cost_in _ _ _ H) = snd x -> P +| None => Z = None -> P +end -> P + + with + | Some x => _ + | None => _ + end _). + +simpl. intros. apply X with (H0 := H). reflexivity. +auto. +Qed. + +Lemma spill_inv : forall g res, +spill g = Some res -> +lowest_cost (get_spillWL (irc_wl g)) (irc_g g) = Some (fst res). +Proof. + intros. + apply spill_inv_aux with g. + rewrite H. + auto. +Qed. + +Lemma spill_inv2 : forall g res, +spill g = Some res -> +exists H, snd res = spill_irc (fst res) g (lowest_cost_in (fst res) (get_spillWL (irc_wl g)) (irc_g g) H). + +Proof. +intros. +apply (spill_inv_aux g). rewrite H. +simpl. intros. rewrite <-H1. +exists H0. reflexivity. +Qed. diff --git a/backend/IRC_graph.v b/backend/IRC_graph.v new file mode 100755 index 00000000..31e9ce80 --- /dev/null +++ b/backend/IRC_graph.v @@ -0,0 +1,15 @@ +Require Import FSets. +Require Import InterfGraphMapImp. + +Record irc_graph := Make_IRC_Graph { +irc_g : tt; +irc_wl : WS; +pal : VertexSet.t; +irc_k : nat; +HWS_irc : WS_properties irc_g irc_k irc_wl; +Hk : VertexSet.cardinal pal = irc_k +}. + +Definition graph_to_IRC_graph g palette := +let K := VertexSet.cardinal palette in +Make_IRC_Graph g (get_WL g K) palette K (WS_prop_get _ _) (refl_equal _). diff --git a/backend/IRC_termination.v b/backend/IRC_termination.v new file mode 100755 index 00000000..f2f7d9f3 --- /dev/null +++ b/backend/IRC_termination.v @@ -0,0 +1,397 @@ +Require Import FSets. +Require Import InterfGraphMapImp. +Require Import IRC_Graph_Functions. +Require Import ZArith. +Require Import IRC_Graph_Functions. +Require Import WS. +Require Import IRC_graph. +Require Import Simplify_WL. +Require Import Spill_WL. +Require Import Freeze_WL. +Require Import Merge_WL. +Require Import MyRegisters. +Require Import Remove_Vertex_WL. +Require Import Affinity_relation. +Require Import Edges. +Require Import Merge_Move. +Require Import Merge_Degree. +Require Import Conservative_criteria. +Require Import Delete_Preference_Edges_Degree. +Require Import Delete_Preference_Edges_Move. +Require Import Remove_Vertex_Move. +Require Import Remove_Vertex_Degree. + +Import Edge RegFacts Props OTFacts. + +Definition not_precolored_irc g := not_precolored (irc_g g). + +Definition irc_measure g := +let simplifyWL := get_simplifyWL (irc_wl g) in +let m := VertexSet.cardinal simplifyWL in +let np := not_precolored_irc g in +let n := VertexSet.cardinal np in +2*n - m. + +Lemma empty_union_empty : +VertexSet.Equal +(VertexSet.union VertexSet.empty VertexSet.empty) +(VertexSet.empty). + +Proof. +split;intros. +destruct (VertexSet.union_1 H); elim (VertexSet.empty_1 H0). +elim (VertexSet.empty_1 H). +Qed. + +Lemma cardinal_union_not_precolored : forall g, +VertexSet.cardinal (not_precolored_irc g) = +VertexSet.cardinal (get_simplifyWL (irc_wl g)) + +VertexSet.cardinal (get_freezeWL (irc_wl g)) + +VertexSet.cardinal (get_spillWL (irc_wl g)). + +Proof. +intros. unfold not_precolored_irc. rewrite <-(not_precolored_union_ws _ _ (irc_wl g)). +do 2 rewrite union_cardinal_inter. rewrite union_inter_1. +generalize (WS_empty_inter_1 _ _ (irc_wl g) (HWS_irc g)). intro. +generalize (WS_empty_inter_2 _ _ (irc_wl g) (HWS_irc g)). intro. +generalize (WS_empty_inter_3 _ _ (irc_wl g) (HWS_irc g)). intro. +generalize (empty_is_empty_1 H). intro. +generalize (empty_is_empty_1 H0). intro. +generalize (empty_is_empty_1 H1). intro. +rewrite H2. rewrite H3. rewrite H4. rewrite empty_union_empty. +rewrite empty_cardinal. omega. +apply HWS_irc. +Qed. + +Lemma irc_measure_equiv_aux : forall g, +irc_measure g = VertexSet.cardinal (get_simplifyWL (irc_wl g)) + + VertexSet.cardinal (get_freezeWL (irc_wl g)) + + VertexSet.cardinal (get_freezeWL (irc_wl g)) + + VertexSet.cardinal (get_spillWL (irc_wl g)) + + VertexSet.cardinal (get_spillWL (irc_wl g)). + +Proof. +intros. unfold irc_measure. simpl. +rewrite cardinal_union_not_precolored. repeat rewrite <-plus_n_O. omega. +Qed. + +Definition card_n g := VertexSet.cardinal (not_precolored_irc g). +Definition card_p g := +VertexSet.cardinal (VertexSet.union (get_freezeWL (irc_wl g)) (get_spillWL (irc_wl g))). + +Lemma irc_measure_equiv : forall g, +irc_measure g = card_n g + card_p g. + +Proof. +intros. unfold card_n, card_p. +rewrite irc_measure_equiv_aux. rewrite cardinal_union_not_precolored. +repeat rewrite <-plus_assoc. +rewrite plus_comm. rewrite <-plus_assoc. +rewrite union_cardinal. omega. +intros. intro. destruct H. elim (WS_empty_inter_1 _ _ _ (HWS_irc g) x). +apply VertexSet.inter_3; auto. +Qed. + +Lemma remove_cardinal_weak_dec : forall x s, +VertexSet.cardinal s - 1 <= VertexSet.cardinal (VertexSet.remove x s). + +Proof. +intros. +destruct (In_dec x s); intros. +rewrite <-(remove_cardinal_1 i). omega. +rewrite <-(remove_cardinal_2 n). omega. +Qed. + +Lemma simplify_dec_aux_1 : forall g r H, +card_n (simplify_irc r g H) < card_n g. + +Proof. +intros. unfold card_n, simplify_irc, not_precolored_irc, not_precolored; simpl. +generalize (In_simplify_props _ _ _ _ _ _ _ _ H (refl_equal _) (HWS_irc g)). intro. +destruct H0. destruct H1. destruct H2. +apply subset_cardinal_lt with (x:=r). +unfold VertexSet.Subset. intros. +generalize (VertexSet.diff_1 H4). intro. +generalize (VertexSet.diff_2 H4). clear H4. intro. +apply VertexSet.diff_3. apply (VertexSet.remove_3 H5). +rewrite precolored_remove_vertex in H4. intro. elim H4. +apply VertexSet.remove_2; auto. +intro. elim (VertexSet.remove_1 H7 H5). +apply VertexSet.diff_3; auto. +intro. generalize (VertexSet.diff_1 H4). intro. + generalize (VertexSet.diff_2 H4). intro. +elim (VertexSet.remove_1 (MyRegisters.Regs.eq_refl _) H5). +Qed. + +Lemma simplify_dec_aux_2 : forall g r H, +card_p (simplify_irc r g H) <= card_p g. + +Proof. +intros. unfold card_p, simplify_irc, not_precolored_irc, not_precolored; simpl. +apply subset_cardinal. unfold simplify_wl. +case_eq (VertexSet.partition (move_related (irc_g g)) + (VertexSet.filter + (fun x : VertexSet.elt => + eq_K (irc_k g) + (VertexSet.cardinal (interference_adj x (irc_g g)))) + (VertexSet.diff (interference_adj r (irc_g g)) + (precolored (irc_g g))))); intros. +unfold get_spillWL, get_freezeWL; simpl. +assert (VertexSet.Equal t0 (fst (VertexSet.partition (move_related (irc_g g)) + (VertexSet.filter + (fun x : VertexSet.elt => + eq_K (irc_k g) + (VertexSet.cardinal (interference_adj x (irc_g g)))) + (VertexSet.diff (interference_adj r (irc_g g)) + (precolored (irc_g g))))))) as Ht0. +rewrite H0. simpl. apply VertexSet.eq_refl. +rewrite VertexSet.partition_1 in Ht0. + +unfold VertexSet.Subset; intros. +destruct (VertexSet.union_1 H1). +destruct (VertexSet.union_1 H2). +apply VertexSet.union_2. auto. +apply VertexSet.union_3. WS_apply (HWS_irc g). rewrite Ht0 in H3. clear H0 H1. +generalize (VertexSet.filter_1 (compat_bool_move _ ) H3). intro. +generalize (VertexSet.filter_2 (compat_bool_move _) H3). clear H3. intro. +generalize (VertexSet.filter_1 (eq_K_compat _ _) H0). intro. +generalize (VertexSet.filter_2 (eq_K_compat _ _) H0). clear H0. intro. +generalize (eq_K_2 _ _ H0). clear H0. intro. +generalize (VertexSet.diff_1 H3). intro. +generalize (VertexSet.diff_2 H3). clear H3. intro. +split. +unfold has_low_degree, interf_degree. rewrite <-H0. +destruct (le_lt_dec (irc_k g) (irc_k g)). reflexivity. elim (lt_irrefl _ l). +split. +apply move_related_in_graph; auto. +assumption. +apply VertexSet.union_3. apply (VertexSet.diff_1 H2). +apply compat_bool_move. +Qed. + +Lemma simplify_dec : forall (g : irc_graph) (r : Register.t) (g' : irc_graph), +simplify g = Some (r, g') -> irc_measure g' < irc_measure g. + +Proof. +intros. do 2 rewrite irc_measure_equiv. +generalize (simplify_inv _ _ H). intro. +generalize (simplify_inv2 _ _ H). intro. simpl in *. +destruct H1 as [H1 H2]. clear H H0. rewrite H2 in *. +generalize (simplify_dec_aux_1 g r (VertexSet.choose_1 H1)). +generalize (simplify_dec_aux_2 g r (VertexSet.choose_1 H1)). omega. +Qed. + +Lemma coalesce_dec_aux_1 : forall g e p q, +~VertexSet.In (snd_ext e) (precolored (irc_g g)) -> +card_n (merge_irc e g p q) < card_n g. + +Proof. +intros. unfold card_n, merge_irc, not_precolored_irc, not_precolored; simpl. +apply subset_cardinal_lt with (x:=(snd_ext e)). +unfold VertexSet.Subset. intros. +generalize (VertexSet.diff_1 H0). intro. +generalize (VertexSet.diff_2 H0). clear H0. intro. +apply VertexSet.diff_3. +apply (VertexSet.remove_3 H1). +rewrite precolored_merge in H0. intro. elim H0. +apply VertexSet.remove_2; auto. +intro. elim (VertexSet.remove_1 H3 H1). +apply VertexSet.diff_3. apply (proj2 (In_graph_edge_in_ext _ _ p)). auto. +intro. generalize (VertexSet.diff_1 H0). intro. + generalize (VertexSet.diff_2 H0). intro. +elim (VertexSet.remove_1 (Regs.eq_refl _) H1). +Qed. + +Lemma coalesce_dec_aux_2 : forall g e p q, +~VertexSet.In (snd_ext e) (precolored (irc_g g)) -> +card_p (merge_irc e g p q) <= card_p g. + +Proof. +intros. unfold card_p, merge_irc, not_precolored_irc, not_precolored; simpl. +apply subset_cardinal. unfold VertexSet.Subset; intros. +destruct (VertexSet.union_1 H0). +generalize (In_freeze_props _ _ _ _ _ _ _ _ H1 (refl_equal _) (ws_merge3 _ _ _ _)). intro. +destruct H2. destruct H3. destruct H4. +case_eq (has_low_degree (irc_g g) (irc_k g) a); intros. +apply VertexSet.union_2. WS_apply (HWS_irc g). +split. assumption. +split. eapply move_related_merge_move_related; eauto. +intro. elim H5. rewrite precolored_merge. apply VertexSet.remove_2. +rewrite In_merge_vertex in H4. destruct H4. auto. +assumption. +apply VertexSet.union_3. WS_apply (HWS_irc g). +split. assumption. +split. rewrite In_merge_vertex in H4. destruct H4. auto. +intro. elim H5. rewrite precolored_merge. apply VertexSet.remove_2. +rewrite In_merge_vertex in H4. destruct H4. auto. +assumption. + +generalize (In_spill_props _ _ _ _ _ _ _ _ H1 (refl_equal _) (ws_merge3 _ _ _ _)). intro. +destruct H2. destruct H3. +rewrite In_merge_vertex in H3. destruct H3. +destruct (Register.eq_dec a (fst_ext e)). +rewrite e0. case_eq (has_low_degree (irc_g g) (irc_k g) (fst_ext e)); intros. +apply VertexSet.union_2. WS_apply (HWS_irc g). +split. assumption. +split. apply (proj1 (Aff_edge_aff _ _ p q)). rewrite e0 in H4. +intro. elim H4. rewrite precolored_merge. apply VertexSet.remove_2. +apply (In_graph_edge_diff_ext _ _ p). +assumption. +apply VertexSet.union_3. WS_apply (HWS_irc g). +split. assumption. +split. rewrite <-e0. assumption. +rewrite e0 in H4. +intro. elim H4. rewrite precolored_merge. apply VertexSet.remove_2. +apply (In_graph_edge_diff_ext _ _ p). +assumption. + +apply VertexSet.union_3. WS_apply (HWS_irc g). +split. rewrite Hk in H2. eapply merge_low_1; eauto. +split. assumption. +intro. elim H4. rewrite precolored_merge. apply VertexSet.remove_2; auto. +Qed. + +Lemma coalesce_dec : forall (g : irc_graph) (e : Edge.t) (g' : irc_graph), +coalesce g = Some (e, g') -> irc_measure g' < irc_measure g. + +Proof. +intros. do 2 rewrite irc_measure_equiv. +generalize (coalesce_inv _ _ H). intro. +generalize (coalesce_inv_2 _ _ H). intro. simpl in *. +destruct H1 as [H1 H2]. destruct H2 as [H2 H3]. clear H. rewrite H3 in *. +cut (forall e', EdgeSet.In e' (get_movesWL (irc_wl g)) -> In_graph_edge e' (irc_g g)). intro HH. +generalize (any_coalescible_edge_1 _ _ _ _ HH H0). intro. destruct H. +generalize (any_coalescible_edge_2 _ _ _ _ HH H0). intro. +generalize (coalesce_dec_aux_1 g e H1 H2 H5). +generalize (coalesce_dec_aux_2 g e H1 H2 H5). omega. +intros. apply (proj2(In_move_props _ _ _ _ _ _ _ _ H (refl_equal _) (HWS_irc g))). +Qed. + +Lemma freeze_dec_aux_1 : forall g x H H0, +card_n (delete_preference_edges_irc2 x g H H0) <= card_n g. + +Proof. +intros. unfold card_n, delete_preference_edges_irc2, not_precolored_irc, not_precolored; simpl. +apply subset_cardinal. unfold VertexSet.Subset; intros. +generalize (VertexSet.diff_1 H1). intro. +generalize (VertexSet.diff_2 H1). clear H1. intro. +rewrite precolored_delete_preference_edges in H1. +apply VertexSet.diff_3; assumption. +Qed. + +Lemma freeze_dec_aux_2 : forall g x H H0, +card_p (delete_preference_edges_irc2 x g H H0) < card_p g. + +Proof. +intros. unfold card_p, delete_preference_edges_irc2, not_precolored_irc, not_precolored; simpl. +generalize (In_freeze_props _ _ _ _ _ _ _ _ H0 (refl_equal _) (HWS_irc g)). intro. +destruct H1. destruct H2. destruct H3. +apply subset_cardinal_lt with (x:=x). unfold VertexSet.Subset; intros. +destruct (VertexSet.union_1 H5). +apply VertexSet.union_2. WS_apply (HWS_irc g). +generalize (In_freeze_props _ _ _ _ _ _ _ _ H6 (refl_equal _) (WS_freeze _ _ H H0)). intro. +destruct H7. destruct H8. destruct H9. +split. +rewrite <-delete_preference_edges_low in H7. auto. +split. eapply move_related_delete_move_related; eauto. +rewrite precolored_delete_preference_edges in H10. assumption. +apply VertexSet.union_3. WS_apply (HWS_irc g). +generalize (In_spill_props _ _ _ _ _ _ _ _ H6 (refl_equal _) (WS_freeze _ _ H H0)). intro. +destruct H7. destruct H8. +split. +rewrite <-delete_preference_edges_low in H7. auto. +split. rewrite In_delete_preference_edges_vertex in H8. auto. +rewrite precolored_delete_preference_edges in H9. assumption. + +apply VertexSet.union_2. assumption. + +intro. +destruct (VertexSet.union_1 H5). +generalize (In_freeze_props _ _ _ _ _ _ _ _ H6 (refl_equal _) (WS_freeze _ _ H H0)). intro. +destruct H7. destruct H8. +rewrite (not_aff_related_delete_preference_edges _ _ H) in H8. congruence. +generalize (In_spill_props _ _ _ _ _ _ _ _ H6 (refl_equal _) (WS_freeze _ _ H H0)). intro. +destruct H7. +rewrite <-delete_preference_edges_low in H7. congruence. +Qed. + +Lemma freeze_dec : forall (g : irc_graph) (r : Register.t) (g' : irc_graph), +freeze g = Some (r, g') -> irc_measure g' < irc_measure g. + +Proof. +intros. do 2 rewrite irc_measure_equiv. +generalize (freeze_inv _ _ H). intro. +generalize (freeze_inv2 _ _ H). intro. simpl in *. +destruct H1 as [H1 H2]. destruct H2 as [H2 H3]. clear H. rewrite H3 in *. +generalize (freeze_dec_aux_1 g r H2 H1). +generalize (freeze_dec_aux_2 g r H2 H1). omega. +Qed. + +Lemma spill_dec_aux_1 : forall g r H, +card_n (spill_irc r g H) < card_n g. + +Proof. +intros. unfold card_n, spill_irc, not_precolored_irc, not_precolored; simpl. +generalize (In_spill_props _ _ _ _ _ _ _ _ H (refl_equal _) (HWS_irc g)). intro. +destruct H0. destruct H1. generalize H2. intro H3. +apply subset_cardinal_lt with (x:=r). +unfold VertexSet.Subset. intros. +generalize (VertexSet.diff_1 H4). intro. +generalize (VertexSet.diff_2 H4). clear H4. intro. +apply VertexSet.diff_3. +apply (VertexSet.remove_3 H5). +rewrite precolored_remove_vertex in H4. intro. elim H4. +apply VertexSet.remove_2; auto. +intro. elim (VertexSet.remove_1 H7 H5). +apply VertexSet.diff_3; auto. +intro. generalize (VertexSet.diff_1 H4). intro. + generalize (VertexSet.diff_2 H4). intro. +elim (VertexSet.remove_1 (Regs.eq_refl _) H5). +Qed. + +Lemma spill_dec_aux_2 : forall g r H, +card_p (spill_irc r g H) <= card_p g. + +Proof. +intros. unfold card_p, spill_irc, not_precolored_irc, not_precolored; simpl. +apply subset_cardinal. unfold VertexSet.Subset; intros. +destruct (VertexSet.union_1 H0). +case_eq (has_low_degree (irc_g g) (irc_k g) a); intros. +apply VertexSet.union_2. WS_apply (HWS_irc g). +generalize (In_freeze_props _ _ _ _ _ _ _ _ H1 (refl_equal _) (WS_spill _ _ H)). intro. +destruct H3. destruct H4. destruct H5. +split. assumption. +split. apply move_remove_2 with (r:=r). assumption. +intro. elim H6. rewrite precolored_remove_vertex. +apply VertexSet.remove_2. rewrite In_remove_vertex in H5. destruct H5. auto. +assumption. +apply VertexSet.union_3. WS_apply (HWS_irc g). +generalize (In_freeze_props _ _ _ _ _ _ _ _ H1 (refl_equal _) (WS_spill _ _ H)). intro. +destruct H3. destruct H4. destruct H5. +split. assumption. +split. rewrite In_remove_vertex in H5. destruct H5. auto. +intro. elim H6. rewrite precolored_remove_vertex. +apply VertexSet.remove_2. rewrite In_remove_vertex in H5. destruct H5. auto. +assumption. +apply VertexSet.union_3. +WS_apply (HWS_irc g). +generalize (In_spill_props _ _ _ _ _ _ _ _ H1 (refl_equal _) (WS_spill _ _ H)). intro. +destruct H2. destruct H3. +rewrite In_remove_vertex in H3. destruct H3. +split. apply not_low_remove_not_low with (r:=r); assumption. +split. assumption. +intro. elim H4. rewrite precolored_remove_vertex. +apply VertexSet.remove_2. auto. assumption. +Qed. + +Lemma spill_dec : forall (g : irc_graph) (r : Register.t) (g' : irc_graph), +spill g = Some (r, g') -> irc_measure g' < irc_measure g. + +Proof. +intros. do 2 rewrite irc_measure_equiv. +generalize (spill_inv _ _ H). intro. +generalize (spill_inv2 _ _ H). intro. simpl in *. +destruct H1 as [H1 H2]. clear H H0. rewrite H2 in *. +generalize (spill_dec_aux_1 g r (lowest_cost_in _ _ _ H1)). +generalize (spill_dec_aux_2 g r (lowest_cost_in _ _ _ H1)). omega. +Qed. diff --git a/backend/InterfGraph.v b/backend/InterfGraph.v index 8a9dda67..081ef59d 100644 --- a/backend/InterfGraph.v +++ b/backend/InterfGraph.v @@ -19,6 +19,11 @@ Require Import Maps. Require Import Ordered. Require Import Registers. Require Import Locations. +Require Import AST. +Require Import Op. +Require Import RTLtyping. +Require Import RTL. +Require Import Conventions. (** Interference graphs are undirected graphs with two kinds of nodes: - RTL pseudo-registers; @@ -298,4 +303,3 @@ Proof. intros. unfold all_interf_regs. apply in_setregreg_fold'. eapply in_setregmreg_fold. eauto. Qed. - diff --git a/backend/InterfGraphMapImp.v b/backend/InterfGraphMapImp.v new file mode 100755 index 00000000..cdd83278 --- /dev/null +++ b/backend/InterfGraphMapImp.v @@ -0,0 +1,9401 @@ +Require Import FSets. +Require Import Recdef. +Require Import ZArith. +Require Import Coq.Init.Wf. +Require Import FSetInterface. +Require Import SetsFacts. +Require Import FMaps. +Require Import OrderedOption. +Require Import FMapAVL. +Require Import Edges. +Require Import MyRegisters. + +Module Register := Regs. + +Import Edge. + +Module VertexSet := FSetAVL.Make Vertex. +Module EdgeSet := FSetAVL.Make Edge. +Module VertexMap := FMapAVL.Make Vertex. +Module MapFacts := Facts VertexMap. +Module RegFacts := MyFacts VertexSet. +Module MEdgeFacts := MyFacts EdgeSet. +Module RegRegProps := MEdgeFacts.Props. +Module Props := RegFacts.Props. + +Definition adj_set x m := +match (VertexMap.find x m) with +| None => VertexSet.empty +| Some x => x +end. + +Record tt : Type := Make_Graph { +vertices : VertexSet.t; +imap : VertexMap.t VertexSet.t; +pmap : VertexMap.t VertexSet.t; +extremities_imap : forall x, VertexMap.In x imap <-> VertexSet.In x vertices; +extremities_pmap : forall x, VertexMap.In x pmap <-> VertexSet.In x vertices; +simple_graph : forall x y, VertexSet.In x (adj_set y imap) /\ + VertexSet.In x (adj_set y pmap) -> False; +sym_imap : forall x y, VertexSet.In x (adj_set y imap) -> + VertexSet.In y (adj_set x imap); +sym_pmap : forall x y, VertexSet.In x (adj_set y pmap) -> + VertexSet.In y (adj_set x pmap); +not_eq_extremities : forall x y, VertexSet.In x (adj_set y imap) \/ + VertexSet.In x (adj_set y pmap) -> + ~Vertex.eq x y +}. + +Definition t := tt. + +Definition V := vertices. + +Definition edgemap_to_edgeset map w := +VertexMap.fold + (fun y imapy s => VertexSet.fold + (fun z s' => EdgeSet.add (y,z,w) s') + imapy + s) + map + EdgeSet.empty. + +Definition AE g := edgemap_to_edgeset (pmap g) (Some N0). + +Definition IE g := edgemap_to_edgeset (imap g) None. + +Definition is_precolored v (g : t) := is_mreg v. + +Definition imap_remove m x := +VertexMap.remove x +(VertexSet.fold + (fun y => VertexMap.add y (VertexSet.remove x (adj_set y m))) + (adj_set x m) + m +). + +Lemma change_fst : forall x y z, +fst_ext (x,y,z) = x. + +Proof. +auto. +Qed. + +Lemma change_snd : forall x y z, +snd_ext (x,y,z) = y. + +Proof. +auto. +Qed. + +Lemma change_weight : forall x y z, +get_weight (x,y,z) = z. + +Proof. +auto. +Qed. + +(* rewriting tactic *) + +Ltac change_rewrite := +repeat (try rewrite change_fst in *;try rewrite change_snd in *;try rewrite change_weight in *). + +(* two tactics for proving equality of edges *) + +Ltac Eq_eq := +apply eq_ordered_eq;unfold E.eq;split;[simpl;split;intuition|simpl;apply OptionN_as_OT.eq_refl]. + +Ltac Eq_comm_eq := rewrite edge_comm;Eq_eq. + +Lemma eq_charac : forall x y, +eq x y -> (Vertex.eq (fst_ext x) (fst_ext y) /\ Vertex.eq (snd_ext x) (snd_ext y)) \/ + (Vertex.eq (fst_ext x) (snd_ext y) /\ Vertex.eq (snd_ext x) (fst_ext y)). + +Proof. +intros x y H;unfold eq in H;unfold ordered_edge in H; +unfold get_edge in H. +destruct (OTFacts.lt_dec (snd_ext x) (fst_ext x)); +destruct (OTFacts.lt_dec (snd_ext y) (fst_ext y)); +unfold E.eq in H;simpl in H;intuition. +Qed. + +Section fold_assoc_interf_map. + +Variable A : Type. + +Inductive eq_set_option : option VertexSet.t -> option VertexSet.t -> Prop := +|None_eq : eq_set_option None None +|Some_eq : forall s s', VertexSet.Equal s s' -> eq_set_option (Some s) (Some s'). + +Definition EqualSetMap m1 m2 := forall x, eq_set_option (VertexMap.find x m1) (VertexMap.find x m2). + +Lemma EqualSetMap_refl : forall m, EqualSetMap m m. + +Proof. +unfold EqualSetMap. intro m. intro x. +destruct (VertexMap.find x m). +constructor. intuition. +constructor. +Qed. + +Lemma EqualSetMap_trans : forall m1 m2 m3, +EqualSetMap m1 m2 -> +EqualSetMap m2 m3 -> +EqualSetMap m1 m3. + +Proof. +intros m1 m2 m3 H H0. +unfold EqualSetMap in *. +intro x. +generalize (H x). clear H. intro H. +generalize (H0 x). clear H0. intro H0. +destruct (VertexMap.find x m1). +inversion H. subst. +rewrite <-H2 in H0. +destruct (VertexMap.find x m3). +constructor. inversion H0. subst. +rewrite H3. assumption. +inversion H0. +destruct (VertexMap.find x m3). +inversion H0. inversion H. subst. rewrite <-H4 in H1. inversion H1. +constructor. +Qed. + +Lemma fold_left_compat_map : forall (f : VertexMap.t VertexSet.t -> A -> VertexMap.t VertexSet.t) l e e', +EqualSetMap e e' -> +(forall e1 e2 a, EqualSetMap e1 e2 -> EqualSetMap (f e1 a) (f e2 a)) -> +EqualSetMap (fold_left f l e) (fold_left f l e'). + +Proof. +intro f;induction l;simpl. +auto. +intros e e' H H0 H1. +apply (IHl (f e a) (f e' a)). +apply H0;assumption. +assumption. +Qed. + +Lemma fold_left_assoc_map : forall l (f : VertexMap.t VertexSet.t -> A -> VertexMap.t VertexSet.t) x h, +(forall (y z : A) s, EqualSetMap (f (f s y) z) (f (f s z) y)) -> +(forall e1 e2 a, EqualSetMap e1 e2 -> EqualSetMap (f e1 a) (f e2 a)) -> +EqualSetMap (fold_left f (h :: l) x) (f (fold_left f l x) h). + +Proof. +induction l;simpl;intros f x h H H0. +apply EqualSetMap_refl. +apply EqualSetMap_trans with (m2 := fold_left f (h :: l) (f x a)). +simpl. apply fold_left_compat_map. apply H. +assumption. +apply IHl. assumption. assumption. +Qed. + +End fold_assoc_interf_map. + +Lemma fold_assoc : forall g g' y0 z s, +(forall x y a, EdgeSet.Equal (g x (g' y a)) (g' y (g x a))) -> +(forall (y z0 : VertexSet.elt) (s0 : EdgeSet.t), +EdgeSet.Equal (g z0 (g y s0)) (g y (g z0 s0))) -> +(forall (e1 e2 : EdgeSet.t) (a1 : VertexSet.elt), +EdgeSet.Equal e1 e2 -> EdgeSet.Equal (g a1 e1) (g a1 e2)) -> +(forall (e1 e2 : EdgeSet.t) (a1 : VertexSet.elt), +EdgeSet.Equal e1 e2 -> EdgeSet.Equal (g' a1 e1) (g' a1 e2)) -> +(forall (y z0 : VertexSet.elt) (s0 : EdgeSet.t), +EdgeSet.Equal (g' z0 (g' y s0)) (g' y (g' z0 s0))) -> +EdgeSet.Equal (VertexSet.fold g z (VertexSet.fold g' y0 s)) + (VertexSet.fold g' y0 (VertexSet.fold g z s)). + +Proof. +intros. +repeat rewrite VertexSet.fold_1. +set (f1 := fun (a : EdgeSet.t) (e : VertexSet.elt) => g e a). +set (f2 := fun (a : EdgeSet.t) (e : VertexSet.elt) => g' e a). +induction (VertexSet.elements z). simpl. +apply EdgeSet.eq_refl. + +set (l' := VertexSet.elements y0) in *. +assert (EdgeSet.Equal (fold_left f2 l' (fold_left f1 (a :: l) s)) + (fold_left f2 l' (f1 (fold_left f1 l s) a))). +apply MEdgeFacts.fold_left_compat_set. +apply MEdgeFacts.fold_left_assoc. + +unfold f2. assumption. +unfold f2. assumption. +unfold f1. assumption. + +apply EdgeSet.eq_trans with (s' := (fold_left f2 l' (f1 (fold_left f1 l s) a))). +set (s' := fold_left f1 l s) in *. +cut (EdgeSet.Equal (fold_left f2 l' (f1 s' a)) (f1 (fold_left f2 l' s') a)). +intro. +apply EdgeSet.eq_trans with (s' := f1 (fold_left f2 l' s') a). +rewrite MEdgeFacts.fold_left_assoc. +apply H1. assumption. +assumption. +assumption. +apply EdgeSet.eq_sym. auto. + +clear IHl. clear H4. +induction l'. simpl. apply EdgeSet.eq_refl. +assert (EdgeSet.Equal (f1 (fold_left f2 (a0 :: l') s') a) + (f1 (f2 (fold_left f2 l' s') a0) a)). +apply H1. +apply MEdgeFacts.fold_left_assoc. +assumption. +assumption. +apply EdgeSet.eq_trans with (s':= (f1 (f2 (fold_left f2 l' s') a0) a)). +rewrite MEdgeFacts.fold_left_assoc. +apply EdgeSet.eq_trans with (s' := f2 (f1 (fold_left f2 l' s') a) a0). +apply H2. assumption. +unfold f1. unfold f2. +unfold EdgeSet.eq. apply EdgeSet.eq_sym. +apply H. +assumption. +assumption. +apply H1. apply EdgeSet.eq_sym. apply MEdgeFacts.fold_left_assoc. +assumption. +assumption. +apply EdgeSet.eq_sym. auto. +Qed. + +Lemma edgemap_to_edgeset_charac : forall m x y (w : option N), +(forall a b, VertexSet.In a (adj_set b m) -> + VertexSet.In b (adj_set a m)) -> +(EdgeSet.In (x,y,w) (edgemap_to_edgeset m w) <-> VertexSet.In y (adj_set x m)). + +Proof. +intros m x y w Hsym. split; intros. +unfold edgemap_to_edgeset in H. +rewrite VertexMap.fold_1 in H. +generalize VertexMap.elements_2. intro. +generalize (H0 _ m). clear H0. intro HH. +induction (VertexMap.elements m). +simpl in H. +elim (EdgeSet.empty_1 H). +set (f := (fun (a : EdgeSet.t) (p : VertexMap.key * VertexSet.t) => + VertexSet.fold + (fun (z : VertexSet.elt) (s' : EdgeSet.t) => + EdgeSet.add (fst p, z, w) s') (snd p) a)) in *. +case_eq a; intros; subst. +rewrite MEdgeFacts.fold_left_assoc in H. +set (s := fold_left f l EdgeSet.empty) in *. +unfold f in H. simpl in H. +assert (EdgeSet.In (x,y,w) s \/ (VertexSet.In y t0 /\ Vertex.eq k x) \/ + (VertexSet.In x t0 /\ Vertex.eq k y)). +clear IHl. intros. +rewrite VertexSet.fold_1 in H. +generalize VertexSet.elements_2. +intro H0. generalize (H0 t0). clear H0. intro Helt. +induction (VertexSet.elements t0). +simpl in H. left. assumption. +rewrite MEdgeFacts.fold_left_assoc in H. +destruct (proj1 (RegRegProps.Dec.F.add_iff _ _ _) H). +fold (eq (k,a,w) (x,y,w)) in H0. +right. +destruct (eq_charac _ _ H0); destruct H1; change_rewrite. +left. split. +apply Helt. left. apply Vertex.eq_sym. assumption. +assumption. +right. split. +apply Helt. left. apply Vertex.eq_sym. assumption. +assumption. +apply IHl0. assumption. +intros. apply Helt. right. assumption. + +intros. apply RegRegProps.add_add. +intros. apply RegRegProps.Dec.F.add_m. apply eq_refl. assumption. + +destruct H0. +apply IHl. +assumption. +intros. apply HH. auto. +destruct H0. +assert (VertexMap.MapsTo x t0 m). +apply HH. +left. +constructor; simpl; intuition. +generalize (VertexMap.find_1 H1). clear H1. intro H1. +unfold adj_set. rewrite H1. intuition. +apply Hsym. +assert (VertexMap.MapsTo y t0 m). +apply HH. +left. +constructor; simpl; intuition. +generalize (VertexMap.find_1 H1). clear H1. intro H1. +unfold adj_set. rewrite H1. intuition. + +unfold f. +intros. set (g := (fun (z0 : VertexSet.elt) (s' : EdgeSet.t) => + EdgeSet.add (fst z, z0, w) s')). +set (g' := fun (z0 : VertexSet.elt) (s' : EdgeSet.t) => + EdgeSet.add (fst y0, z0, w) s'). +apply fold_assoc. +unfold g. unfold g'. +intros. apply RegRegProps.add_add. +unfold g. unfold g'. +intros. apply RegRegProps.add_add. +intros. apply RegRegProps.Dec.F.add_m. apply eq_refl. assumption. +intros. apply RegRegProps.Dec.F.add_m. apply eq_refl. assumption. +unfold g. unfold g'. +intros. apply RegRegProps.add_add. +intros. +unfold f. +rewrite VertexSet.fold_1. +rewrite VertexSet.fold_1. +apply MEdgeFacts.fold_left_compat_set. +assumption. +intros. apply RegRegProps.Dec.F.add_m. apply eq_refl. assumption. + +unfold edgemap_to_edgeset. +rewrite VertexMap.fold_1. +generalize VertexMap.elements_1. intro. +case_eq (VertexMap.find x m); intros. +generalize (H0 _ m x t0). clear H0. intro HH. +induction (VertexMap.elements m). +simpl. +assert (VertexMap.MapsTo x t0 m). +apply VertexMap.find_2. assumption. +generalize (HH H0). intro H2. inversion H2. + +set (f := (fun (a : EdgeSet.t) (p : VertexMap.key * VertexSet.t) => + VertexSet.fold + (fun (z : VertexSet.elt) (s' : EdgeSet.t) => + EdgeSet.add (fst p, z, w) s') (snd p) a)) in *. +rewrite MEdgeFacts.fold_left_assoc. +set (s := fold_left f l EdgeSet.empty) in *. +unfold f. +destruct a. simpl. +rewrite VertexSet.fold_1. +generalize VertexSet.elements_1. +intro H2. generalize (H2 t1 y). clear H2. intro HHH. +induction (VertexSet.elements t1). simpl. +assert (VertexMap.MapsTo x t0 m). +apply VertexMap.find_2. assumption. +generalize (HH H0). intro H2. +inversion H2; subst. +inversion H4. simpl in H3. simpl in H5. subst. +unfold adj_set in H. rewrite H1 in H. +generalize (HHH H). intro. inversion H5. +apply IHl. intro. auto. + +rewrite MEdgeFacts.fold_left_assoc. +generalize (VertexMap.find_2 H1). intro. +generalize (HH H0). clear HH H0. intro HH. +inversion HH; subst. +inversion H2; simpl in *; subst. clear H2 HH. +destruct (Vertex.eq_dec y a). +apply EdgeSet.add_1. +Eq_eq. +apply EdgeSet.add_2. +apply IHl0. +intro H2. generalize (HHH H2). clear HHH H2. intro H2. +inversion H2. subst. +elim (n H4). +assumption. +apply EdgeSet.add_2. +assert (forall l', EdgeSet.In (x,y,w) s -> + EdgeSet.In (x, y, w) + (fold_left + (fun (a0 : EdgeSet.t) (e : VertexSet.elt) => EdgeSet.add (k, e, w) a0) + l' s)). +clear H HH H1 H2 IHl IHl0 HHH Hsym. +intros. +induction l'. simpl. assumption. +rewrite MEdgeFacts.fold_left_assoc. +destruct (Edge.eq_dec (k,a0,w) (x,y,w)). +apply EdgeSet.add_1. auto. +apply EdgeSet.add_2. apply IHl'. + +intros. apply RegRegProps.add_add. +intros. apply RegRegProps.Dec.F.add_m. apply eq_refl. assumption. + +apply H0. +apply IHl. +auto. + +intros. apply RegRegProps.add_add. +intros. apply RegRegProps.Dec.F.add_m. apply eq_refl. assumption. + +unfold f. +intros. set (g := (fun (z0 : VertexSet.elt) (s' : EdgeSet.t) => + EdgeSet.add (fst z, z0, w) s')). +set (g' := fun (z0 : VertexSet.elt) (s' : EdgeSet.t) => + EdgeSet.add (fst y0, z0, w) s'). +apply fold_assoc. +unfold g. unfold g'. +intros. apply RegRegProps.add_add. +unfold g. unfold g'. +intros. apply RegRegProps.add_add. +intros. apply RegRegProps.Dec.F.add_m. apply eq_refl. assumption. +intros. apply RegRegProps.Dec.F.add_m. apply eq_refl. assumption. +unfold g. unfold g'. +intros. apply RegRegProps.add_add. +intros. +unfold f. +rewrite VertexSet.fold_1. +rewrite VertexSet.fold_1. +apply MEdgeFacts.fold_left_compat_set. +assumption. +intros. apply RegRegProps.Dec.F.add_m. apply eq_refl. assumption. + +unfold adj_set in H. rewrite H1 in H. elim (VertexSet.empty_1 H). +Qed. + +Require Import FMapFacts. + +Module InterfFacts := FMapFacts.Facts VertexMap. + +Lemma imap_remove_1 : forall x y m r, +~Vertex.eq r x -> +~Vertex.eq r y -> +VertexSet.In x (adj_set y m) -> +VertexSet.In x (adj_set y (imap_remove m r)). + +Proof. +intros. +unfold imap_remove. +unfold adj_set. +cut (VertexSet.In x match + (VertexMap.find (elt:=VertexSet.t) y + (VertexMap.remove (elt:=VertexSet.t) r + (VertexSet.fold + (fun y0 : VertexSet.elt => + VertexMap.add y0 (VertexSet.remove r (adj_set y0 m))) + (adj_set r m) m))) with + | Some x0 => x0 + | None => VertexSet.empty +end); auto. +rewrite MapFacts.remove_neq_o. + +rewrite VertexSet.fold_1. +induction (VertexSet.elements (adj_set r m)); intros. +simpl. assumption. + +set (f:= (fun (a : VertexMap.t VertexSet.t) (e : VertexSet.elt) => + VertexMap.add e (VertexSet.remove r (adj_set e m)) a)) in *. +assert (EqualSetMap (fold_left f (a :: l) m) (f (fold_left f l m) a)). +apply fold_left_assoc_map. +intros. +unfold f. unfold EqualSetMap. intros. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +constructor. +apply RegFacts.Props.Dec.F.remove_m. apply Vertex.eq_refl. +unfold adj_set. rewrite (MapFacts.find_o _ (Vertex.eq_trans (Vertex.eq_sym e) e0)). +apply VertexSet.eq_refl. apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply RegFacts.Props.Dec.F.remove_m. apply Vertex.eq_refl. +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply RegFacts.Props.Dec.F.remove_m. apply Vertex.eq_refl. +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. + +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s); constructor; auto. +apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +intros. +unfold f. +unfold EqualSetMap. +intros. +destruct (Vertex.eq_dec x0 a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +apply Vertex.eq_sym. assumption. +apply Vertex.eq_sym. assumption. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H2. auto. auto. +unfold EqualSetMap in H2. +generalize (H2 y). intro H3. +case_eq (VertexMap.find y (fold_left f (a :: l) m)); intros. +rewrite H4 in H3. inversion H3. subst. +unfold f in H6. +destruct (Vertex.eq_dec y a). +rewrite MapFacts.add_eq_o in H6. inversion H6. +rewrite H7. rewrite H8. apply VertexSet.remove_2. +auto. unfold adj_set. rewrite (MapFacts.find_o _ (Vertex.eq_sym e)). assumption. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o in H6. +rewrite H7. unfold f in IHl. +case_eq (VertexMap.find (elt:=VertexSet.t) y + (fold_left + (fun (a : VertexMap.t VertexSet.t) (e : VertexSet.elt) => + VertexMap.add e (VertexSet.remove r (adj_set e m)) a) l m)); intros; +rewrite H5 in *. +inversion H6. subst. +assumption. +inversion H6. +auto. +rewrite H4 in H3. inversion H3. subst. +unfold f in H5. +destruct (Vertex.eq_dec y a). +rewrite MapFacts.add_eq_o in H5. +inversion H5. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o in H5. +unfold f in IHl. +case_eq (VertexMap.find (elt:=VertexSet.t) y + (fold_left + (fun (a : VertexMap.t VertexSet.t) (e : VertexSet.elt) => + VertexMap.add e (VertexSet.remove r (adj_set e m)) a) l m)); intros; +rewrite H6 in *. +inversion H5. +assumption. +auto. +auto. +Qed. + +Lemma imap_remove_2 : forall x y m r, +Vertex.eq r x \/ Vertex.eq r y -> +(forall a b, VertexSet.In a (adj_set b m) -> + VertexSet.In b (adj_set a m)) -> +~VertexSet.In x (adj_set y (imap_remove m r)). + +Proof. +intros x y m r H HH. +unfold imap_remove. +unfold adj_set. +cut (~VertexSet.In x match + (VertexMap.find (elt:=VertexSet.t) y + (VertexMap.remove (elt:=VertexSet.t) r + (VertexSet.fold + (fun y0 : VertexSet.elt => + VertexMap.add y0 (VertexSet.remove r (adj_set y0 m))) + (adj_set r m) m))) with + | Some x0 => x0 + | None => VertexSet.empty +end); auto. +destruct (Vertex.eq_dec r y). +rewrite MapFacts.remove_eq_o. +apply VertexSet.empty_1. +assumption. +rewrite MapFacts.remove_neq_o. +destruct H. + +generalize VertexSet.elements_1. intro HHH. +generalize (HHH (adj_set r m) y). clear HHH. intro HHH. +rewrite VertexSet.fold_1. +induction (VertexSet.elements (adj_set r m)); intros. +simpl. intro H0. +generalize (HH _ _ H0). intro H1. +assert (InA Vertex.eq y nil). +apply HHH. +unfold adj_set. rewrite (MapFacts.find_o _ H). assumption. +inversion H2. + +set (f:= (fun (a : VertexMap.t VertexSet.t) (e : VertexSet.elt) => + VertexMap.add e (VertexSet.remove r (adj_set e m)) a)) in *. +assert (EqualSetMap (fold_left f (a :: l) m) (f (fold_left f l m) a)). +apply fold_left_assoc_map. +intros. +unfold f. unfold EqualSetMap. intros. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +constructor. +apply RegFacts.Props.Dec.F.remove_m. apply Vertex.eq_refl. +unfold adj_set. rewrite (MapFacts.find_o _ (Vertex.eq_trans (Vertex.eq_sym e) e0)). +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply RegFacts.Props.Dec.F.remove_m. apply Vertex.eq_refl. +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply RegFacts.Props.Dec.F.remove_m. apply Vertex.eq_refl. +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s); constructor; auto. +apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +intros. +unfold f. +unfold EqualSetMap. +intros. +destruct (Vertex.eq_dec x0 a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +apply Vertex.eq_sym. assumption. +apply Vertex.eq_sym. assumption. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H0. auto. auto. + +unfold EqualSetMap in H0. +generalize (H0 y). intro H1. +case_eq (VertexMap.find y (fold_left f (a :: l) m)); intros. +rewrite H2 in H1. inversion H1. subst. +unfold f in H4. +destruct (Vertex.eq_dec y a). +rewrite MapFacts.add_eq_o in H4. inversion H4. +rewrite H5. rewrite H6. apply VertexSet.remove_1. auto. apply Regs.eq_sym. auto. + +rewrite MapFacts.add_neq_o in H4. +rewrite H5. unfold f in IHl. +case_eq (VertexMap.find (elt:=VertexSet.t) y + (fold_left + (fun (a : VertexMap.t VertexSet.t) (e : VertexSet.elt) => + VertexMap.add e (VertexSet.remove r (adj_set e m)) a) l m)); intros; +rewrite H3 in *. +inversion H4. subst. +apply IHl. +intros; intuition. +inversion H7; subst. +elim (n0 H9). assumption. +inversion H4. +auto. +apply VertexSet.empty_1. +elim (n H). +auto. +Qed. + +Lemma imap_remove_3 : forall x y m r, +VertexSet.In x (adj_set y (imap_remove m r)) -> +VertexSet.In x (adj_set y m). + +Proof. +intros. +assert (VertexSet.In x match + (VertexMap.find (elt:=VertexSet.t) y + (VertexMap.remove (elt:=VertexSet.t) r + (VertexSet.fold + (fun y0 : VertexSet.elt => + VertexMap.add y0 (VertexSet.remove r (adj_set y0 m))) + (adj_set r m) m))) with + | Some x0 => x0 + | None => VertexSet.empty +end) by auto. generalize H0. clear H H0. intro H. +destruct (Vertex.eq_dec y r). +rewrite MapFacts.remove_eq_o in H. +elim (VertexSet.empty_1 H). apply Regs.eq_sym. auto. +rewrite MapFacts.remove_neq_o in H. +unfold adj_set. +rewrite VertexSet.fold_1 in H. + +generalize VertexSet.elements_2. intro. +generalize (H0 (adj_set y m) x). clear H0. intro HH. + +induction (VertexSet.elements (adj_set r m)). simpl in H. assumption. + +set (f:= (fun (a : VertexMap.t VertexSet.t) (e : VertexSet.elt) => + VertexMap.add e (VertexSet.remove r (adj_set e m)) a)) in *. +assert (EqualSetMap (fold_left f (a :: l) m) (f (fold_left f l m) a)). +apply fold_left_assoc_map. +intros. +unfold f. unfold EqualSetMap. intros. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +constructor. +apply RegFacts.Props.Dec.F.remove_m. apply Vertex.eq_refl. +unfold adj_set. rewrite (MapFacts.find_o _ (Vertex.eq_trans (Vertex.eq_sym e) e0)). +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply RegFacts.Props.Dec.F.remove_m. apply Vertex.eq_refl. +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply RegFacts.Props.Dec.F.remove_m. apply Vertex.eq_refl. +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s); constructor; auto. +apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +intros. +unfold f. +unfold EqualSetMap. +intros. +destruct (Vertex.eq_dec x0 a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +apply Vertex.eq_sym. assumption. +apply Vertex.eq_sym. assumption. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H0. auto. auto. + +unfold EqualSetMap in H0. +generalize (H0 y). +case_eq (VertexMap.find y (fold_left f (a :: l) m)); intros. +rewrite H1 in *. inversion H2. subst. +unfold f in H4. +destruct (Vertex.eq_dec y a). +rewrite MapFacts.add_eq_o in H4. +apply HH. +apply VertexSet.elements_1. +unfold adj_set. rewrite (MapFacts.find_o _ e). fold (adj_set a m). +apply VertexSet.remove_3 with (x:=r). +inversion H4. subst. rewrite <-H5. assumption. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o in H4. +fold f in H4. rewrite <-H4 in IHl. +apply IHl. +rewrite <-H5. assumption. +auto. + +inversion H2. subst. rewrite <-H3 in *. rewrite H1 in *. +elim (VertexSet.empty_1 H). +auto. +Qed. + +Lemma imap_remove_4 : forall x m r, +VertexMap.In x (imap_remove m r) -> +(forall a b, VertexSet.In a (adj_set b m) -> + VertexSet.In b (adj_set a m)) -> +VertexMap.In x m. + +Proof. +intros x m r H Hsym. +unfold imap_remove in H. +destruct (Vertex.eq_dec x r). +elim (VertexMap.remove_1 (Vertex.eq_sym e) H). +apply (proj2 (MapFacts.in_find_iff _ _)). +generalize (proj1 (MapFacts.in_find_iff _ _) H). clear H. intro H. +rewrite MapFacts.remove_neq_o in H. +rewrite VertexSet.fold_1 in H. intro H0. elim H. clear H. + +generalize VertexSet.elements_2. intro. +generalize (H (adj_set r m) x). clear H. intro HH. + +induction (VertexSet.elements (adj_set r m)). simpl. +assumption. + +set (f:= (fun (a : VertexMap.t VertexSet.t) (e : VertexSet.elt) => + VertexMap.add e (VertexSet.remove r (adj_set e m)) a)) in *. +assert (EqualSetMap (fold_left f (a :: l) m) (f (fold_left f l m) a)). +apply fold_left_assoc_map. +intros. +unfold f. unfold EqualSetMap. intros. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y). +rewrite MapFacts.add_eq_o. +constructor. +apply RegFacts.Props.Dec.F.remove_m. apply Vertex.eq_refl. +unfold adj_set. rewrite (MapFacts.find_o _ (Vertex.eq_trans (Vertex.eq_sym e) e0)). +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply RegFacts.Props.Dec.F.remove_m. apply Vertex.eq_refl. +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply RegFacts.Props.Dec.F.remove_m. apply Vertex.eq_refl. +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s); constructor; auto. +apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +intros. +unfold f. +unfold EqualSetMap. +intros. +destruct (Vertex.eq_dec x0 a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +apply Vertex.eq_sym. assumption. +apply Vertex.eq_sym. assumption. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H. auto. auto. + +generalize (H x). clear H. intro H. +case_eq (VertexMap.find x (fold_left f (a :: l) m)); intros. +rewrite H1 in H. inversion H. subst. +unfold f in H3. +destruct (Vertex.eq_dec x a). +assert (VertexSet.In r (adj_set x m)). +apply Hsym. apply HH. left. auto. +unfold adj_set in H2. rewrite H0 in H2. elim (VertexSet.empty_1 H2). +rewrite MapFacts.add_neq_o in H3. +fold f in H3. +assert (VertexMap.find x (fold_left f l m) = None). +apply IHl. +intro. apply HH. auto. +rewrite H2 in H3. inversion H3. +auto. +reflexivity. +auto. +Qed. + +Lemma imap_remove_5 : forall r x m, +VertexMap.In x m -> +~Vertex.eq x r -> +VertexMap.In x (imap_remove m r). + +Proof. +intros. +unfold imap_remove. +rewrite MapFacts.in_find_iff. +rewrite MapFacts.remove_neq_o. +rewrite VertexSet.fold_1. +set (f := (fun (a : VertexMap.t VertexSet.t) (e : VertexSet.elt) => + VertexMap.add e (VertexSet.remove r (adj_set e m)) a)). +induction (VertexSet.elements (adj_set r m)). simpl. +rewrite MapFacts.in_find_iff in H. assumption. + +cut (VertexMap.find x (f (fold_left f l m) a) <> None). +intro H1. +cut (EqualSetMap (fold_left f (a :: l) m) (f (fold_left f l m) a)). intro H2. +generalize (H2 x). clear H2. intro H2. inversion H2. +simpl. rewrite <-H4 in *. rewrite <-H5 in *. assumption. +simpl. rewrite <-H3 in *. congruence. + +apply fold_left_assoc_map. +unfold EqualSetMap. unfold f. intros. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y). +rewrite MapFacts.add_eq_o. +constructor. +apply RegFacts.Props.Dec.F.remove_m. apply Vertex.eq_refl. +unfold adj_set. rewrite (MapFacts.find_o _ (Vertex.eq_trans (Vertex.eq_sym e) e0)). +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply RegFacts.Props.Dec.F.remove_m. apply Vertex.eq_refl. +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply RegFacts.Props.Dec.F.remove_m. apply Vertex.eq_refl. +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s); constructor; auto. +apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +intros. +unfold f. +unfold EqualSetMap. +intros. +destruct (Vertex.eq_dec x0 a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +apply Vertex.eq_sym. assumption. +apply Vertex.eq_sym. assumption. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H2. auto. auto. + +set (tmp := fold_left f l m) in *. +unfold f. +destruct (Vertex.eq_dec x a). +rewrite MapFacts.add_eq_o. congruence. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +assumption. +auto. +auto. +Qed. + +Lemma extremities_in_remove_vertex_imap v g : +forall x, +VertexMap.In x (imap_remove (imap g) v) <-> +VertexSet.In x (VertexSet.remove v (V g)). + +Proof. +split; intros. +apply VertexSet.remove_2. +unfold imap_remove in H. +intro Helim; apply (VertexMap.remove_1 Helim H). +apply (extremities_imap g x); apply imap_remove_4 with (r:=v); try assumption. +exact (sym_imap g). + +generalize (proj2 (extremities_imap g x)). intro. +generalize (H0 (VertexSet.remove_3 H)). clear H0. intro H0. +apply imap_remove_5. assumption. +intro H1. elim (VertexSet.remove_1 (Vertex.eq_sym H1) H). +Qed. + +Lemma extremities_in_remove_vertex_pmap v g : +forall x, +VertexMap.In x (imap_remove (pmap g) v) <-> +VertexSet.In x (VertexSet.remove v (V g)). + +Proof. +split; intros. +apply VertexSet.remove_2. +unfold imap_remove in H. +intro Helim; apply (VertexMap.remove_1 Helim H). +apply (extremities_pmap g x); apply imap_remove_4 with (r:=v); try assumption. +exact (sym_pmap g). + +generalize (proj2 (extremities_pmap g x)). intro. +generalize (H0 (VertexSet.remove_3 H)). clear H0. intro H0. +apply imap_remove_5. assumption. +intro H1. elim (VertexSet.remove_1 (Vertex.eq_sym H1) H). +Qed. + +Lemma simple_graph_remove_vertex_map v g : +forall x y, +VertexSet.In x (adj_set y (imap_remove (imap g) v)) /\ +VertexSet.In x (adj_set y (imap_remove (pmap g) v)) -> +False. + +Proof. +intros. +apply (simple_graph g x y). +destruct H. +generalize (imap_remove_3 _ _ _ _ H). +generalize (imap_remove_3 _ _ _ _ H0). +auto. +Qed. + +Lemma not_eq_extremities_remove_vertex_map v g : forall x y, +VertexSet.In x (adj_set y (imap_remove (imap g) v)) \/ +VertexSet.In x (adj_set y (imap_remove (pmap g) v)) -> +~Vertex.eq x y. + +Proof. +intros. +apply (not_eq_extremities g). +destruct H;[left|right]; apply (imap_remove_3 _ _ _ _ H). +Qed. + +Lemma sym_imap_remove_vertex v g : +forall (x : VertexSet.elt) (y : VertexMap.key), +VertexSet.In x (adj_set y (imap_remove (imap g) v)) -> +VertexSet.In y (adj_set x (imap_remove (imap g) v)). + +Proof. +intros. +generalize (imap_remove_2 x y (imap g) v). intro H0. +apply imap_remove_1. +intro H1. elim (H0 (or_intror _ H1) (sym_imap g) H). +intro H1. elim (H0 (or_introl _ H1) (sym_imap g) H). +generalize (imap_remove_3 _ _ _ _ H). intro H1. +apply (sym_imap g). assumption. +Qed. + +Lemma sym_pmap_remove_vertex v g : +forall (x : VertexSet.elt) (y : VertexMap.key), +VertexSet.In x (adj_set y (imap_remove (pmap g) v)) -> +VertexSet.In y (adj_set x (imap_remove (pmap g) v)). + +Proof. +intros. +generalize (imap_remove_2 x y (pmap g) v). intro H0. +apply imap_remove_1. +intro H1. elim (H0 (or_intror _ H1) (sym_pmap g) H). +intro H1. elim (H0 (or_introl _ H1) (sym_pmap g) H). +generalize (imap_remove_3 _ _ _ _ H). intro H1. +apply (sym_pmap g). assumption. +Qed. + +Definition remove_vertex v g := +Make_Graph (VertexSet.remove v (V g)) + (imap_remove (imap g) v) + (imap_remove (pmap g) v) + (extremities_in_remove_vertex_imap v g) + (extremities_in_remove_vertex_pmap v g) + (simple_graph_remove_vertex_map v g) + (sym_imap_remove_vertex v g) + (sym_pmap_remove_vertex v g) + (not_eq_extremities_remove_vertex_map v g). + +Definition map_merge e map := +let adj_snd := VertexSet.remove (fst_ext e) (adj_set (snd_ext e) map) in +let adj_fst := VertexSet.remove (snd_ext e) (adj_set (fst_ext e) map) in +let new_fst_adj := VertexSet.union adj_fst adj_snd in +(* +let new_fst_adj_ := VertexSet.union (adj_set (fst_ext e) map) (adj_set (snd_ext e) map) in +let new_fst_adj := VertexSet.remove (fst_ext e) (VertexSet.remove (snd_ext e) new_fst_adj_) in +let m := VertexMap.add (fst_ext e) new_fst_adj map in +*) +let redirect_m := + VertexSet.fold + (fun y m' => + VertexMap.add y + (VertexSet.add (fst_ext e) (VertexSet.remove (snd_ext e) (adj_set y map))) m') + adj_snd +(* + (VertexSet.remove (fst_ext e) (adj_set (snd_ext e) m)) +*) + map in +VertexMap.remove (snd_ext e) +(VertexMap.add (fst_ext e) new_fst_adj redirect_m). + +Definition imap_merge e g := map_merge e (imap g). + +Definition resolve_conflicts y pm padj iadj := +let m' := VertexSet.fold + (fun x m => VertexMap.add x (VertexSet.remove y (adj_set x pm)) m) + (VertexSet.inter padj iadj) + pm +in VertexMap.add y (VertexSet.diff padj iadj) m'. + +Definition pmap_merge e g im := +let pm := map_merge e (pmap g) in +resolve_conflicts (fst_ext e) pm (adj_set (fst_ext e) pm) (adj_set (fst_ext e) im). + +Definition In_graph_edge e g := +EdgeSet.In e (AE g) \/ EdgeSet.In e (IE g). + +Lemma In_graph_edge_dec : forall e g, +{In_graph_edge e g}+{~In_graph_edge e g}. + +Proof. +intros e g. +destruct (RegRegProps.In_dec e (AE g)). +left. left. assumption. +destruct (RegRegProps.In_dec e (IE g)). +left. right. assumption. +right. intro H. destruct H;[elim n|elim n0];assumption. +Qed. + +Lemma aff_edge_dec : forall e, +{aff_edge e}+{~aff_edge e}. + +Proof. +intro e. +case_eq (get_weight e). +intros n H. left. unfold aff_edge. exists n. auto. +intro H. right. intro H0. unfold aff_edge in H0. +destruct H0 as [w H0]. rewrite H0 in H. inversion H. +Qed. + +Definition In_graph (v : Vertex.t) g := VertexSet.In v (V g). + +Add Morphism get_weight : get_weight_m. + +Proof. +intros x y H. +rewrite (weight_ordered_weight x);rewrite (weight_ordered_weight y). +unfold get_weight;unfold E.eq in H. +destruct H as [_ H];inversion H;[|rewrite H2];reflexivity. +Qed. + +Lemma E_weights_aux : forall e map w s, +EdgeSet.In e +(VertexMap.fold + (fun y imapy s => VertexSet.fold + (fun z s' => EdgeSet.add (y,z,w) s') + imapy + s) + map + s) -> +EdgeSet.In e s \/ get_weight e = w. + +Proof. +intros. +rewrite VertexMap.fold_1 in H. +generalize H. clear H. generalize (VertexMap.elements map) s. +induction l. +simpl. auto. +intros. simpl in H. +set (s' := VertexSet.fold + (fun (z : VertexSet.elt) (s' : EdgeSet.t) => + EdgeSet.add (fst a, z, w) s') (snd a) s0) in H. +generalize (IHl s' H). intro H0. +destruct H0. +unfold s' in H0. + +assert (EdgeSet.In e (VertexSet.fold (fun (z : VertexSet.elt) (s' : EdgeSet.t) => + EdgeSet.add (fst a, z, w) s') (snd a) s0) -> EdgeSet.In e s0 \/ get_weight e = w). +clear H IHl. +rewrite VertexSet.fold_1. +induction (VertexSet.elements (snd a)). +simpl. auto. +rewrite MEdgeFacts.fold_left_assoc. +intro H1. destruct (proj1 (RegRegProps.Dec.F.add_iff _ _ _) H1). +fold (eq (fst a, a0, w) e) in H. +right. rewrite <-H. simpl. reflexivity. +apply IHl0. assumption. + +intros. apply RegRegProps.add_add. +intros. apply RegRegProps.Dec.F.add_m. +apply eq_refl. +assumption. +apply H1. assumption. +right. assumption. +Qed. + +Lemma E_weights : forall e m w, +EdgeSet.In e (edgemap_to_edgeset m w) -> get_weight e = w. + +Proof. +intros. +generalize (E_weights_aux e m w EdgeSet.empty). intro H0. +destruct H0. +assumption. +elim (EdgeSet.empty_1 H0). +assumption. +Qed. + +Lemma IE_weights : forall g e, +EdgeSet.In e (IE g) -> get_weight e = None. + +Proof. +unfold IE. intros. eapply E_weights. eassumption. +Qed. + +Lemma AE_weights : forall g e, +EdgeSet.In e (AE g) -> get_weight e = Some N0. + +Proof. +unfold AE. intros. eapply E_weights. eassumption. +Qed. + +(* extremities of edges are in the graph *) +Lemma In_graph_edge_in_ext : forall e g, +In_graph_edge e g -> In_graph (fst_ext e) g /\ In_graph (snd_ext e) g. + +Proof. +intros. +split. destruct H. +apply (proj1 (extremities_pmap g (fst_ext e))). + +generalize (AE_weights _ _ H). intro Hw. +unfold AE in *. +rewrite (edge_eq e) in H. +simpl in Hw. rewrite Hw in H. +generalize (proj1 (edgemap_to_edgeset_charac _ _ _ _(sym_pmap g)) H). intro H0. +apply (proj2 (InterfFacts.in_find_iff _ _)). +unfold adj_set in H0. +case_eq (VertexMap.find (fst_ext e) (pmap g)); intros; rewrite H1 in H0. +intro Helim. inversion Helim. +elim (VertexSet.empty_1 H0). + +apply (proj1 (extremities_imap g (fst_ext e))). +generalize (IE_weights _ _ H). intro Hw. +unfold IE in *. +rewrite (edge_eq e) in H. +simpl in Hw. rewrite Hw in H. +generalize (proj1 (edgemap_to_edgeset_charac _ _ _ _ (sym_imap g)) H). intro H0. +apply (proj2 (InterfFacts.in_find_iff _ _)). +unfold adj_set in H0. +case_eq (VertexMap.find (fst_ext e) (imap g)); intros; rewrite H1 in H0. +intro Helim. inversion Helim. +elim (VertexSet.empty_1 H0). + +destruct H. +apply (proj1 (extremities_pmap g (snd_ext e))). +generalize (AE_weights _ _ H). intro Hw. +unfold AE in *. +rewrite (edge_eq e) in H. rewrite edge_comm in H. +simpl in Hw. rewrite Hw in H. +generalize (proj1 (edgemap_to_edgeset_charac _ _ _ _ (sym_pmap g)) H). intro H0. +apply (proj2 (InterfFacts.in_find_iff _ _)). +unfold adj_set in H0. +case_eq (VertexMap.find (snd_ext e) (pmap g)); intros; rewrite H1 in H0. +intro Helim. inversion Helim. +elim (VertexSet.empty_1 H0). + +apply (proj1 (extremities_imap g (snd_ext e))). +generalize (IE_weights _ _ H). intro Hw. +unfold IE in *. +rewrite (edge_eq e) in H. rewrite edge_comm in H. +simpl in Hw. rewrite Hw in H. +generalize (proj1 (edgemap_to_edgeset_charac _ _ _ _ (sym_imap g)) H). intro H0. +apply (proj2 (InterfFacts.in_find_iff _ _)). +unfold adj_set in H0. +case_eq (VertexMap.find (snd_ext e) (imap g)); intros; rewrite H1 in H0. +intro Helim. inversion Helim. +elim (VertexSet.empty_1 H0). +Qed. + +Lemma not_eq_extremities_map_merge : forall x y e m, +(forall a b, VertexSet.In a (adj_set b m) -> ~Vertex.eq a b) -> +VertexSet.In x (adj_set y (map_merge e m)) -> +~Vertex.eq x y. + +Proof. +intros x y e m Hsimp H. +unfold map_merge in H. +set (f := (fun (y : VertexSet.elt) (m' : VertexMap.t VertexSet.t) => + VertexMap.add y + (VertexSet.add (fst_ext e) + (VertexSet.remove (snd_ext e) (adj_set y m))) m')) in *. +set (s := (VertexSet.union + (VertexSet.remove (snd_ext e) (adj_set (fst_ext e) m)) + (VertexSet.remove (fst_ext e) (adj_set (snd_ext e) m)))) in *. +set (s' := (VertexSet.remove (fst_ext e) (adj_set (snd_ext e) m))) in *. +intro. +unfold adj_set in H. rewrite MapFacts.remove_neq_o in H. +destruct (Vertex.eq_dec y (fst_ext e)). +rewrite MapFacts.add_eq_o in H. +unfold s in H. +destruct (VertexSet.union_1 H). +generalize (VertexSet.remove_3 H1). clear H1. intro H1. +rewrite H0 in H1. rewrite e0 in H1. +elim (Hsimp _ _ H1). auto. +unfold s' in H1. +rewrite H0 in H1. rewrite e0 in H1. elim (VertexSet.remove_1 (Vertex.eq_refl _) H1). +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o in H. + +rewrite VertexSet.fold_1 in H. +set (f' := fun a e => f e a) in *. +induction (VertexSet.elements s'). simpl in H. +unfold adj_set in H. +fold (adj_set y m) in H. +elim (Hsimp _ _ H). assumption. + +cut (EqualSetMap (fold_left f' (a :: l) m) (f' (fold_left f' l m) a)). intro. +generalize (H1 y). clear H1. intro H1. simpl in H. inversion H1; clear H1. +unfold adj_set in H. rewrite <-H3 in H. elim (VertexSet.empty_1 H). + +rewrite <-H2 in H. +set (tmp := fold_left f' l m) in *. +unfold f' in H3. unfold f in H3. +destruct (Vertex.eq_dec y a). +rewrite MapFacts.add_eq_o in H3. +inversion H3. subst. clear H3. +rewrite H4 in H. clear H4. +rewrite H0 in H. +destruct (proj1 (Props.Dec.F.add_iff _ _ _) H). +elim (n (Vertex.eq_sym H1)). +generalize (VertexSet.remove_3 H1). intro. +elim (Hsimp _ _ H3). auto. apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o in H3. +rewrite <-H3 in IHl. +apply IHl. rewrite <-H4. assumption. +auto. + +apply fold_left_assoc_map. +unfold EqualSetMap. intros. +unfold f'. unfold f. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +constructor. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s0); constructor; apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f'. unfold f. +destruct (Vertex.eq_dec x0 a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H1. auto. +auto. + +auto. + +intro. rewrite MapFacts.remove_eq_o in H. elim (VertexSet.empty_1 H). auto. +Qed. + +Lemma resolve_conflicts_map_0 : forall x y e g, +VertexSet.In x (adj_set y (resolve_conflicts (fst_ext e) (map_merge e (pmap g)) + (adj_set (fst_ext e) (map_merge e (pmap g))) + (adj_set (fst_ext e) (imap_merge e g)))) -> +VertexSet.In x (adj_set y (map_merge e (pmap g))). + +Proof. +intros. +unfold resolve_conflicts in H. +set (f := (fun (x : VertexSet.elt) (m : VertexMap.t VertexSet.t) => + VertexMap.add x (VertexSet.remove (fst_ext e) (adj_set x (map_merge e (pmap g)))) m)) in *. +set (s := (VertexSet.diff (adj_set (fst_ext e) (map_merge e (pmap g))) + (adj_set (fst_ext e) (imap_merge e g)))) in *. +set (inter := (VertexSet.inter (adj_set (fst_ext e) (map_merge e (pmap g))) + (adj_set (fst_ext e) (imap_merge e g)))) in *. +unfold adj_set in H. +destruct (Vertex.eq_dec y (fst_ext e)). rewrite MapFacts.add_eq_o in H. +unfold s in H. +generalize (VertexSet.diff_1 H). intro H0. +unfold adj_set. rewrite (MapFacts.find_o _ e0). assumption. +apply Regs.eq_sym. auto. + +rewrite MapFacts.add_neq_o in H. +rewrite VertexSet.fold_1 in H. +set (f' := fun a e => f e a) in *. +set (m := map_merge e (pmap g)) in *. +induction (VertexSet.elements inter). simpl in H. assumption. +cut (EqualSetMap (fold_left f' (a :: l) m) (f' (fold_left f' l m) a)). intro. +generalize (H0 y). clear H0. simpl in H. intro H0. inversion H0; clear H0. +rewrite <-H2 in H. elim (VertexSet.empty_1 H). +set (tmp := fold_left f' l m) in *. +unfold f' in H2. unfold f in H2. +destruct (Vertex.eq_dec y a). rewrite MapFacts.add_eq_o in H2. +rewrite <-H1 in H. rewrite H3 in H. clear H3. +inversion H2; subst; clear H2. +unfold adj_set. rewrite (MapFacts.find_o _ e0). +apply (VertexSet.remove_3 H). +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o in H2. +apply IHl. rewrite <-H2. rewrite <-H1 in H. rewrite <-H3. assumption. +auto. + +apply fold_left_assoc_map. +unfold EqualSetMap. intros. +unfold f'. unfold f. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +constructor. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s0); constructor; apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f'. unfold f. +destruct (Vertex.eq_dec x0 a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H0. auto. +auto. +auto. +Qed. + +Lemma pmap_merge_sub : forall x y e g, +VertexSet.In x (adj_set y (pmap_merge e g (imap_merge e g))) -> +VertexSet.In x (adj_set y (map_merge e (pmap g))). + +Proof. +exact (fun x y e g H => resolve_conflicts_map_0 _ _ _ _ H). +Qed. + +Lemma pmap_merge_domain_1 : forall x e g, +In_graph_edge e g -> +VertexMap.In x (map_merge e (pmap g)) -> +VertexSet.In x (VertexSet.remove (snd_ext e) (V g)). + +Proof. +intros x e g HH H. +unfold map_merge in H. +set (m := pmap g) in *. +set (s := VertexSet.remove (fst_ext e) (adj_set (snd_ext e) m)) in *. +set (adj := (VertexSet.union + (VertexSet.remove (snd_ext e) (adj_set (fst_ext e) m)) s)) in *. +set (f := (fun (y : VertexSet.elt) (m'0 : VertexMap.t VertexSet.t) => + VertexMap.add y + (VertexSet.add (fst_ext e) + (VertexSet.remove (snd_ext e) (adj_set y m))) m'0)) in *. +rewrite VertexSet.fold_1 in H. +set (f' := fun a e => f e a) in *. +rewrite MapFacts.in_find_iff in H. +assert (~Vertex.eq x (snd_ext e)). +intro. rewrite MapFacts.remove_eq_o in H. congruence. apply Regs.eq_sym. auto. +apply VertexSet.remove_2. auto. +rewrite MapFacts.remove_neq_o in H. +destruct (Vertex.eq_dec x (fst_ext e)). +rewrite e0. apply (proj1 (In_graph_edge_in_ext _ _ HH)). + +cut (forall z, VertexSet.In z s -> VertexSet.In z (V g)). intro HHH. +generalize VertexSet.elements_2. intro H1. +generalize (H1 s x). clear H1. intro Helt. + +induction (VertexSet.elements s). simpl in H. +apply (proj1 (extremities_pmap g x)). +rewrite MapFacts.in_find_iff. +rewrite MapFacts.add_neq_o in H. +assumption. +auto. + +cut (EqualSetMap (fold_left f' (a :: l) m) (f' (fold_left f' l m) a)). intro. +generalize (H1 x). clear H1. intro H1. simpl in H. inversion H1. +rewrite MapFacts.add_neq_o in H. +rewrite <-H3 in H. congruence. auto. +set (tmp := fold_left f' l m) in *. +unfold f' in H3. unfold f in H3. +destruct (Vertex.eq_dec x a). +apply HHH. apply Helt. left. auto. +rewrite MapFacts.add_neq_o in H3. +apply IHl. +rewrite MapFacts.add_neq_o. +congruence. +auto. +intro. apply Helt. right. auto. auto. + +apply fold_left_assoc_map. +unfold EqualSetMap. unfold f'. unfold f. intros. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y). +rewrite MapFacts.add_eq_o. +constructor. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s0); constructor; apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. unfold f'. unfold f. intros. +destruct (Vertex.eq_dec x0 a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H1. +auto. +auto. + +intros. +unfold s in H1. +generalize (VertexSet.remove_3 H1). clear H1. intro H1. +apply (proj1 (extremities_pmap g z)). +rewrite MapFacts.in_find_iff. +generalize (sym_pmap g _ _ H1). clear H1. intro H1. +unfold adj_set in H1. +destruct (VertexMap.find z (pmap g)). +congruence. +elim (VertexSet.empty_1 H1). +auto. +Qed. + +Lemma pmap_merge_domain_2 : forall x e g, +In_graph_edge e g -> +VertexSet.In x (VertexSet.remove (snd_ext e) (V g)) -> +VertexMap.In x (map_merge e (pmap g)). + +Proof. +intros x e g HH H. +unfold map_merge. +set (m := pmap g) in *. +set (s := VertexSet.remove (fst_ext e) (adj_set (snd_ext e) m)) in *. +set (adj := (VertexSet.union + (VertexSet.remove (snd_ext e) (adj_set (fst_ext e) m)) s)) in *. +set (f := (fun (y : VertexSet.elt) (m'0 : VertexMap.t VertexSet.t) => + VertexMap.add y + (VertexSet.add (fst_ext e) + (VertexSet.remove (snd_ext e) (adj_set y m))) m'0)) in *. +rewrite VertexSet.fold_1. +set (f' := fun a e => f e a) in *. +rewrite MapFacts.in_find_iff. +assert (~Vertex.eq x (snd_ext e)). +intro. elim (VertexSet.remove_1 (Vertex.eq_sym H0) H). +generalize (VertexSet.remove_3 H). clear H. intro H. +rewrite MapFacts.remove_neq_o. +destruct (Vertex.eq_dec x (fst_ext e)). +rewrite MapFacts.add_eq_o. congruence. apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. + +induction (VertexSet.elements s). simpl. +rewrite <-MapFacts.in_find_iff. +apply (proj2 (extremities_pmap g x)). assumption. + +cut (EqualSetMap (fold_left f' (a :: l) m) (f' (fold_left f' l m) a)). intro. +generalize (H1 x). clear H1. intro H1. simpl. inversion H1. +set (tmp := fold_left f' l m) in *. +unfold f' in H4. unfold f in H4. +destruct (Vertex.eq_dec x a). +rewrite MapFacts.add_eq_o in H4. congruence. apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o in H4. congruence. auto. +congruence. + +apply fold_left_assoc_map. +unfold EqualSetMap. unfold f'. unfold f. intros. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y). +rewrite MapFacts.add_eq_o. +constructor. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s0); constructor; apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. unfold f'. unfold f. intros. +destruct (Vertex.eq_dec x0 a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H1. +auto. +auto. + +auto. +auto. +Qed. + +(* there is no loop in the graph *) +Lemma In_graph_edge_diff_ext : forall e g, +In_graph_edge e g -> ~Vertex.eq (snd_ext e) (fst_ext e). + +Proof. +intros. +apply (not_eq_extremities g). +destruct H;[right|left]. + +generalize (AE_weights _ _ H). intro Hw. +unfold AE in *. +rewrite (edge_eq e) in H. +simpl in Hw. rewrite Hw in H. +apply (proj1 (edgemap_to_edgeset_charac _ _ _ _(sym_pmap g)) H). + +generalize (IE_weights _ _ H). intro Hw. +unfold IE in *. +rewrite (edge_eq e) in H. +simpl in Hw. rewrite Hw in H. +apply (proj1 (edgemap_to_edgeset_charac _ _ _ _(sym_imap g)) H). +Qed. + +Lemma extremities_in_merge_map e g : +In_graph_edge e g -> +(forall x, +VertexMap.In x (map_merge e (pmap g)) <-> +VertexSet.In x (VertexSet.remove (snd_ext e) (V g))). + +Proof. +intros e g HH x; split; intro H0. +generalize H0. intro H. +unfold map_merge in H0. +apply VertexSet.remove_2. +intro H1. +elim (VertexMap.remove_1 H1 H0). +generalize (proj1 (MapFacts.remove_in_iff _ _ _) H0). clear H0. +intro H0. destruct H0. +generalize (proj1 (MapFacts.add_in_iff _ _ _ _) H1). clear H1. +intro H1. +destruct H1. +rewrite <-H1. +apply (proj1 (In_graph_edge_in_ext _ _ HH)). +apply (proj1 (extremities_pmap g x)). +rewrite VertexSet.fold_1 in H1. +set (f := (fun (a : VertexMap.t VertexSet.t) (e0 : VertexSet.elt) => + VertexMap.add e0 + (VertexSet.add (fst_ext e) + (VertexSet.remove (snd_ext e) (adj_set e0 (pmap g)))) a)) in *. +set (s := VertexSet.remove (fst_ext e) (adj_set (snd_ext e) (pmap g))) in *. +cut (forall z, In z (VertexSet.elements s) -> + VertexMap.In z (pmap g)). +intros HHH. +induction (VertexSet.elements s). +simpl in H1. assumption. +set (m := pmap g) in *. +cut (VertexMap.In x (f (fold_left f l m) a)). +clear H1. intro H1. +set (tmp := fold_left f l m) in *. +unfold f in H1. +destruct (Vertex.eq_dec x a). +rewrite e0. apply (HHH a). left. auto. +rewrite MapFacts.in_find_iff in H1. +rewrite MapFacts.add_neq_o in H1. +apply IHl. +rewrite MapFacts.in_find_iff. +assumption. +intros. apply HHH. right. assumption. +auto. + +assert (EqualSetMap (fold_left f (a :: l) m) (f (fold_left f l m) a)). +apply fold_left_assoc_map. + +unfold EqualSetMap. intros. +unfold f. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y). +rewrite MapFacts.add_eq_o. +constructor. +apply RegFacts.Props.Dec.F.add_m. +apply Regs.eq_refl. +apply RegFacts.Props.Dec.F.remove_m. +apply Regs.eq_refl. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +rewrite (MapFacts.find_o _ e1). +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply RegFacts.Props.Dec.F.add_m. +apply Regs.eq_refl. +apply RegFacts.Props.Dec.F.remove_m. +apply Regs.eq_refl. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s0); constructor. apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f. +destruct (Vertex.eq_dec x0 a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H2. auto. +auto. +rewrite MapFacts.in_find_iff. +generalize (H2 x). intro. +inversion H3. +rewrite MapFacts.in_find_iff in H1. +simpl in H1. rewrite <-H5 in H1. assumption. +congruence. + +intros. +exists (adj_set z (pmap g)). apply VertexMap.elements_2. +apply VertexMap.elements_1. +unfold adj_set. +case_eq (VertexMap.find z (pmap g)); intros. +apply VertexMap.find_2. assumption. + +assert (VertexSet.In z s). +apply VertexSet.elements_2. +apply In_InA. apply Regs.eq_refl. assumption. +unfold s in H4. +generalize (VertexSet.remove_3 H4). clear H4. intro H4. +generalize (sym_pmap g _ _ H4). clear H4. intro H4. +unfold adj_set in H4. +rewrite H3 in H4. elim (VertexSet.empty_1 H4). + +unfold map_merge. +rewrite MapFacts.in_find_iff. +rewrite MapFacts.remove_neq_o. +destruct (Vertex.eq_dec x (fst_ext e)). +rewrite MapFacts.add_eq_o. +congruence. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite VertexSet.fold_1. +set (f := (fun (a : VertexMap.t VertexSet.t) (e0 : VertexSet.elt) => + VertexMap.add e0 + (VertexSet.add (fst_ext e) + (VertexSet.remove (snd_ext e) (adj_set e0 (pmap g)))) a)) in *. +set (s := VertexSet.remove (fst_ext e) (adj_set (snd_ext e) (pmap g))) in *. +induction (VertexSet.elements s). simpl. +generalize (proj2 (extremities_pmap g x) (VertexSet.remove_3 H0)). +rewrite MapFacts.in_find_iff. auto. + +set (m := pmap g) in *. +assert (EqualSetMap (fold_left f (a :: l) m) (f (fold_left f l m) a)). +apply fold_left_assoc_map. + +unfold EqualSetMap. intros. +unfold f. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y). +rewrite MapFacts.add_eq_o. +constructor. +apply RegFacts.Props.Dec.F.add_m. +apply Regs.eq_sym. auto. +apply RegFacts.Props.Dec.F.remove_m. +apply Regs.eq_sym. auto. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +rewrite (MapFacts.find_o _ e1). +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply RegFacts.Props.Dec.F.add_m. +apply Regs.eq_sym. auto. +apply RegFacts.Props.Dec.F.remove_m. +apply Regs.eq_refl. +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s0); constructor. apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f. +destruct (Vertex.eq_dec x0 a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +apply Regs.eq_sym. auto. +apply Regs.eq_sym. auto. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H. auto. +auto. + +generalize (H x). clear H. intro H. +inversion H. +simpl. +set (tmp := fold_left f l m) in *. +unfold f in H3. +destruct (Vertex.eq_dec x a). +rewrite MapFacts.add_eq_o in H3. +congruence. +intuition. +rewrite MapFacts.add_neq_o in H3. +rewrite H3 in IHl. congruence. +auto. +simpl. rewrite <-H1. congruence. +auto. +intro. elim (VertexSet.remove_1 H H0). +Qed. + +Lemma extremities_in_merge_imap e g : +In_graph_edge e g -> +(forall x, +VertexMap.In x (imap_merge e g) <-> +VertexSet.In x (VertexSet.remove (snd_ext e) (V g))). + +Proof. +intros e g HH x; split; intro H0. +generalize H0. intro H. +unfold imap_merge in H0. +unfold map_merge in H0. +apply VertexSet.remove_2. +intro H1. +elim (VertexMap.remove_1 H1 H0). +generalize (proj1 (MapFacts.remove_in_iff _ _ _) H0). clear H0. +intro H0. destruct H0. +generalize (proj1 (MapFacts.add_in_iff _ _ _ _) H1). clear H1. +intro H1. +destruct H1. +rewrite <-H1. +apply (proj1 (In_graph_edge_in_ext _ _ HH)). +apply (proj1 (extremities_imap g x)). +rewrite VertexSet.fold_1 in H1. +set (f := (fun (a : VertexMap.t VertexSet.t) (e0 : VertexSet.elt) => + VertexMap.add e0 + (VertexSet.add (fst_ext e) + (VertexSet.remove (snd_ext e) (adj_set e0 (imap g)))) a)) in *. +set (s := VertexSet.remove (fst_ext e) (adj_set (snd_ext e) (imap g))) in *. +cut (forall z, In z (VertexSet.elements s) -> + VertexMap.In z (imap g)). +intros HHH. +induction (VertexSet.elements s). +simpl in H1. assumption. +set (m := imap g) in *. +cut (VertexMap.In x (f (fold_left f l m) a)). +clear H1. intro H1. +set (tmp := fold_left f l m) in *. +unfold f in H1. +destruct (Vertex.eq_dec x a). +rewrite e0. apply (HHH a). left. auto. +rewrite MapFacts.in_find_iff in H1. +rewrite MapFacts.add_neq_o in H1. +apply IHl. +rewrite MapFacts.in_find_iff. +assumption. +intros. apply HHH. right. assumption. +auto. + +assert (EqualSetMap (fold_left f (a :: l) m) (f (fold_left f l m) a)). +apply fold_left_assoc_map. + +unfold EqualSetMap. intros. +unfold f. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y). +rewrite MapFacts.add_eq_o. +constructor. +apply RegFacts.Props.Dec.F.add_m. +intuition. +apply RegFacts.Props.Dec.F.remove_m. +intuition. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +rewrite (MapFacts.find_o _ e1). +apply VertexSet.eq_refl. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply RegFacts.Props.Dec.F.add_m. +intuition. +apply RegFacts.Props.Dec.F.remove_m. +intuition. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s0); constructor. apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f. +destruct (Vertex.eq_dec x0 a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H2. auto. +auto. +rewrite MapFacts.in_find_iff. +generalize (H2 x). intro. +inversion H3. +rewrite MapFacts.in_find_iff in H1. +simpl in H1. rewrite <-H5 in H1. assumption. +congruence. + +intros. +exists (adj_set z (imap g)). apply VertexMap.elements_2. +apply VertexMap.elements_1. +unfold adj_set. +case_eq (VertexMap.find z (imap g)); intros. +apply VertexMap.find_2. assumption. + +assert (VertexSet.In z s). +apply VertexSet.elements_2. +apply In_InA. intuition. auto. +unfold s in H4. +generalize (VertexSet.remove_3 H4). clear H4. intro H4. +generalize (sym_imap g _ _ H4). clear H4. intro H4. +unfold adj_set in H4. +rewrite H3 in H4. elim (VertexSet.empty_1 H4). + +unfold imap_merge. +unfold map_merge. +rewrite MapFacts.in_find_iff. +rewrite MapFacts.remove_neq_o. +destruct (Vertex.eq_dec x (fst_ext e)). +rewrite MapFacts.add_eq_o. +congruence. +intuition. +rewrite MapFacts.add_neq_o. +rewrite VertexSet.fold_1. +set (f := (fun (a : VertexMap.t VertexSet.t) (e0 : VertexSet.elt) => + VertexMap.add e0 + (VertexSet.add (fst_ext e) + (VertexSet.remove (snd_ext e) (adj_set e0 (imap g)))) a)) in *. +set (s := VertexSet.remove (fst_ext e) (adj_set (snd_ext e) (imap g))) in *. +induction (VertexSet.elements s). simpl. +generalize (proj2 (extremities_imap g x) (VertexSet.remove_3 H0)). +rewrite MapFacts.in_find_iff. auto. + +set (m := imap g) in *. +assert (EqualSetMap (fold_left f (a :: l) m) (f (fold_left f l m) a)). +apply fold_left_assoc_map. + +unfold EqualSetMap. intros. +unfold f. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y). +rewrite MapFacts.add_eq_o. +constructor. +apply RegFacts.Props.Dec.F.add_m. +intuition. +apply RegFacts.Props.Dec.F.remove_m. +intuition. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +rewrite (MapFacts.find_o _ e1). +apply VertexSet.eq_refl. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply RegFacts.Props.Dec.F.add_m. +intuition. +apply RegFacts.Props.Dec.F.remove_m. +intuition. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s0); constructor. apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f. +destruct (Vertex.eq_dec x0 a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H. auto. +auto. + +generalize (H x). clear H. intro H. +inversion H. +simpl. +set (tmp := fold_left f l m) in *. +unfold f in H3. +destruct (Vertex.eq_dec x a). +rewrite MapFacts.add_eq_o in H3. +congruence. +intuition. +rewrite MapFacts.add_neq_o in H3. +rewrite H3 in IHl. congruence. +intuition. +simpl. rewrite <-H1. congruence. +auto. +intro. elim (VertexSet.remove_1 H H0). +Qed. + +Lemma sym_map_merge_map : forall e g m x y, +aff_edge e -> +In_graph_edge e g -> +(forall a b, VertexSet.In a (adj_set b m) -> + VertexSet.In b (adj_set a m)) -> +(forall a b, VertexSet.In a (adj_set b (map_merge e m)) -> + ~Vertex.eq a b) -> +~Vertex.eq x (snd_ext e) -> +~Vertex.eq y (snd_ext e) -> +VertexSet.In x (adj_set y (map_merge e m)) -> +VertexSet.In y (adj_set x (map_merge e m)). + +Proof. +intros e g m x y Haff Hin Hsym Hnoteq Hsndx Hsndy H. +assert (~Vertex.eq x y) as Hdiff. +apply Hnoteq. assumption. +unfold map_merge in *. +set (f := (fun (y : VertexSet.elt) (m' : VertexMap.t VertexSet.t) => + VertexMap.add y + (VertexSet.add (fst_ext e) + (VertexSet.remove (snd_ext e) (adj_set y m))) + m')) in *. +set (s := (VertexSet.union + (VertexSet.remove (snd_ext e) + (adj_set (fst_ext e) m)) + (VertexSet.remove (fst_ext e) + (adj_set (snd_ext e) m)))) in *. +set (s' := adj_set (snd_ext e) m) in *. +cut (forall z, VertexSet.In z s' -> ~Vertex.eq z (fst_ext e) -> VertexSet.In z s). intro Himp. +unfold adj_set. rewrite MapFacts.remove_neq_o. +destruct (Vertex.eq_dec x (fst_ext e)). +rewrite MapFacts.add_eq_o. + +unfold adj_set in H. +rewrite MapFacts.remove_neq_o in H. +destruct (Vertex.eq_dec y (fst_ext e)). +rewrite MapFacts.add_eq_o in H. +rewrite e1. rewrite <-e0. assumption. +intuition. + +rewrite MapFacts.add_neq_o in H. +case_eq (VertexMap.find y (VertexSet.fold f (VertexSet.remove (fst_ext e) s') m)); intros. +rewrite H0 in H. + +rewrite VertexSet.fold_1 in H0. +set (f' := fun a e => f e a) in *. + +assert (eq_set_option (VertexMap.find (elt:=VertexSet.t) y + (fold_left f' (VertexSet.elements (VertexSet.remove (fst_ext e) s')) + m)) (Some t0)). rewrite H0. constructor. apply VertexSet.eq_refl. +generalize H1. clear H0 H1. intro H0. +generalize VertexSet.elements_2. intro. +generalize (H1 (VertexSet.remove (fst_ext e) s')). clear H1. intro HH. +induction (VertexSet.elements (VertexSet.remove (fst_ext e) s')). +unfold s. apply VertexSet.union_2. +apply VertexSet.remove_2. +auto. +apply Hsym. rewrite <-e0. +unfold adj_set. inversion H0. subst. rewrite H3. assumption. + +cut (EqualSetMap (fold_left f' (a :: l) m) (f' (fold_left f' l m) a)). intro. +generalize (H1 y). clear H1. intro H1. simpl in H0. inversion H1. +inversion H0. congruence. +set (tmp := fold_left f' l m) in *. +unfold f' in H3. unfold f in H3. +destruct (Vertex.eq_dec y a). +rewrite e1. +apply Himp. +apply VertexSet.remove_3 with (x:= fst_ext e). +apply HH. left. intuition. +assert (VertexSet.In a (VertexSet.remove (fst_ext e) s')). +apply HH. left. intuition. +intro. elim (VertexSet.remove_1 (Vertex.eq_sym H6) H5). +rewrite MapFacts.add_neq_o in H3. clear H1. +apply IHl. +inversion H0. subst. +rewrite <-H1 in H2. clear H1. inversion H2. subst. clear H2. +rewrite <-H3. constructor. +rewrite <-H4. assumption. + +intros. apply HH. right. assumption. +auto. + +apply fold_left_assoc_map. +unfold EqualSetMap. unfold f'. unfold f. +intros. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +constructor. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). +rewrite (MapFacts.find_o _ (Vertex.eq_sym e2)). +apply VertexSet.eq_refl. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s0); constructor; apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f'. unfold f. +destruct (Vertex.eq_dec x0 a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H1. auto. auto. + +rewrite H0 in H. +elim (VertexSet.empty_1 H). +auto. +auto. +intuition. + +rewrite MapFacts.add_neq_o. +unfold adj_set in H. +rewrite MapFacts.remove_neq_o in H. +destruct (Vertex.eq_dec y (fst_ext e)). +rewrite MapFacts.add_eq_o in H. + +rewrite VertexSet.fold_1. +generalize VertexSet.elements_1. intro H0. +generalize (H0 (VertexSet.remove (fst_ext e) s') x). clear H0. intro H0. +induction (VertexSet.elements (VertexSet.remove (fst_ext e) s')). simpl. +unfold s in H. +destruct (VertexSet.union_1 H). +generalize (VertexSet.remove_3 H1). clear H1. intro H1. +generalize (Hsym _ _ H1). clear H1. intro H1. +rewrite <-e0 in H1. assumption. +generalize (VertexSet.remove_3 H1). clear H1. intro H1. +assert (VertexSet.In x (VertexSet.remove (fst_ext e) s')) by + (apply VertexSet.remove_2; auto). +generalize (H0 H2). intro H3. inversion H3. +set (f' := fun a e => f e a) in *. + +cut (EqualSetMap (fold_left f' (a :: l) m) (f' (fold_left f' l m) a)). intro. +generalize (H1 x). clear H1. intro H1. simpl. inversion H1. +set (tmp := fold_left f' l m) in *. +unfold f' in H4. unfold f in H4. +destruct (Vertex.eq_dec x a). +rewrite MapFacts.add_eq_o in H4. congruence. +intuition. +rewrite MapFacts.add_neq_o in H4. rewrite <-H4 in IHl. +apply IHl. intro. +generalize (H0 H2). clear H2. intro H2. +inversion H2; subst. +elim (n0 H6). +assumption. +auto. + +set (tmp := fold_left f' l m) in *. +unfold f' in H3. unfold f in H3. +destruct (Vertex.eq_dec x a). +rewrite MapFacts.add_eq_o in H3. +inversion H3. subst. clear H3. clear H1. +rewrite H4. apply VertexSet.add_1. intuition. intuition. +rewrite MapFacts.add_neq_o in H3. rewrite <-H3 in IHl. +rewrite H4. apply IHl. intro. +generalize (H0 H5). clear H5. intro H5. +inversion H5; subst. +elim (n0 H7). +assumption. +auto. + +apply fold_left_assoc_map. +unfold EqualSetMap. unfold f'. unfold f. +intros. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +constructor. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). +rewrite (MapFacts.find_o _ (Vertex.eq_sym e2)). +apply VertexSet.eq_refl. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s0); constructor; apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f'. unfold f. +destruct (Vertex.eq_dec x0 a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H1. auto. auto. + +intuition. + +rewrite MapFacts.add_neq_o in H. +rewrite VertexSet.fold_1 in H. rewrite VertexSet.fold_1. +set (f' := fun a e => f e a) in *. +induction (VertexSet.elements (VertexSet.remove (fst_ext e) s')). simpl in H. simpl. +apply (Hsym _ _ H). + +cut (EqualSetMap (fold_left f' (a :: l) m) (f' (fold_left f' l m) a)). intro. +generalize (H0 x). generalize (H0 y). clear H0. intros H0 HH0. simpl. simpl in H. inversion H0. +rewrite <-H2 in H. elim (VertexSet.empty_1 H). +set (tmp := fold_left f' l m) in *. +unfold f' in H2. unfold f in H2. +destruct (Vertex.eq_dec y a). +rewrite MapFacts.add_eq_o in H2. +inversion HH0. +unfold f' in H6. unfold f in H6. +destruct (Vertex.eq_dec x a). +elim Hdiff. apply (Vertex.eq_trans e1 (Vertex.eq_sym e0)). +rewrite MapFacts.add_neq_o in H6. + +cut (VertexMap.In x tmp). +intro H7. rewrite MapFacts.in_find_iff in H7. congruence. +cut (VertexMap.In x m -> VertexMap.In x tmp). +intro H7. apply H7. clear H7. +rewrite MapFacts.in_find_iff. intro. + +rewrite <-H1 in *. rewrite H3 in H. clear H0 HH0 H1 H5 H6 IHl. +inversion H2. subst. clear H2. clear H3. +destruct (proj1 (Props.Dec.F.add_iff _ _ _) H). elim n. auto. +generalize (VertexSet.remove_3 H0). clear H H0. intro H. +generalize (Hsym _ _ H). clear H. intro H. +unfold adj_set in H. rewrite H4 in H. elim (VertexSet.empty_1 H). + +clear IHl H H0 HH0 H1 H2 H3 H5 H6. intro. +unfold tmp. +induction l. simpl. assumption. + +cut (EqualSetMap (fold_left f' (a0 :: l) m) (f' (fold_left f' l m) a0)). intro. +generalize (H0 x). clear H0. simpl. intro H0. inversion H0. +set (tmp' := fold_left f' l m) in *. +unfold f' in H3. unfold f in H3. +destruct (Vertex.eq_dec x a0). +rewrite MapFacts.add_eq_o in H3. congruence. intuition. +rewrite MapFacts.add_neq_o in H3. +rewrite MapFacts.in_find_iff in IHl. congruence. intuition. +rewrite MapFacts.in_find_iff. congruence. + +apply fold_left_assoc_map. +unfold EqualSetMap. unfold f'. unfold f. +intros. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +constructor. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). +rewrite (MapFacts.find_o _ (Vertex.eq_sym e2)). +apply VertexSet.eq_refl. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s1); constructor; apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f'. unfold f. +destruct (Vertex.eq_dec x0 a1). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H0. auto. auto. + +auto. +clear H0 HH0. +rewrite <-H1 in H. +unfold f' in H5. unfold f in H5. +destruct (Vertex.eq_dec x a). +elim Hdiff. apply Vertex.eq_trans with (y:=a); auto. +rewrite MapFacts.add_neq_o in H5. +rewrite <-H5 in *. +inversion H2. subst. clear H2. +rewrite H3 in H. +destruct (proj1 (Props.Dec.F.add_iff _ _ _) H). elim n. auto. +generalize (VertexSet.remove_3 H0). clear H H0. intro H. + +case_eq (VertexMap.find y tmp); intros. +rewrite H0 in IHl. +rewrite H6. apply IHl. + +clear IHl H1 H6 H4 H3 H5. +unfold tmp in H0. +unfold adj_set in H. rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)) in H. +assert (eq_set_option (VertexMap.find y (fold_left f' l m)) (Some t0)). +rewrite H0. constructor. apply VertexSet.eq_refl. +generalize H1. clear H0 H1. intro H0. +induction l. simpl in H0. inversion H0. subst. clear H0. rewrite <-H1 in H. rewrite <-H3. assumption. +cut (EqualSetMap (fold_left f' (a0 :: l) m) (f' (fold_left f' l m) a0)). intro. +generalize (H1 y). clear H1. simpl in H0. intro H1. inversion H1. clear H1. +set (tmp' := fold_left f' l m) in *. +unfold f' in H4. unfold f in H4. +destruct (Vertex.eq_dec y a0). +rewrite MapFacts.add_eq_o in H4. congruence. intuition. +rewrite MapFacts.add_neq_o in H4. + +case_eq (VertexMap.find y m); intros. +clear H. + +assert (VertexMap.In y tmp'). +clear H3 H4 IHl H0. +unfold tmp'. induction l. simpl. +rewrite MapFacts.in_find_iff. congruence. + +cut (EqualSetMap (fold_left f' (a1 :: l) m) (f' (fold_left f' l m) a1)). intro. +generalize (H y). clear H. intro H. simpl. inversion H. clear H. +set (tmp'' := fold_left f' l m) in *. +unfold f' in H3. unfold f in H3. +destruct (Vertex.eq_dec y a1). +rewrite MapFacts.add_eq_o in H3. congruence. intuition. +rewrite MapFacts.add_neq_o in H3. +rewrite MapFacts.in_find_iff in IHl. rewrite <-H3 in IHl. congruence. +auto. +rewrite MapFacts.in_find_iff. rewrite <-H0. congruence. + +apply fold_left_assoc_map. +unfold EqualSetMap. unfold f'. unfold f. +intros. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +constructor. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). +rewrite (MapFacts.find_o _ (Vertex.eq_sym e2)). +apply VertexSet.eq_refl. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s2); constructor; apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f'. unfold f. +destruct (Vertex.eq_dec x0 a2). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H. auto. auto. + +rewrite MapFacts.in_find_iff in H. congruence. +rewrite H1 in H. elim (VertexSet.empty_1 H). +auto. + +rewrite <-H2 in *. inversion H0. subst. clear H0 H1. + +set (tmp' := fold_left f' l m) in *. +unfold f' in H3. unfold f in H3. +destruct (Vertex.eq_dec y a0). rewrite MapFacts.add_eq_o in H3. +inversion H3. subst. clear H3. +case_eq (VertexMap.find y m); intros. +rewrite H0 in H. +rewrite <-H7. rewrite H4. +apply VertexSet.add_2. +apply VertexSet.remove_2. +auto. +unfold adj_set. rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). +rewrite H0. assumption. +rewrite H0 in H. elim (VertexSet.empty_1 H). +intuition. +rewrite MapFacts.add_neq_o in H3. + +apply IHl. +rewrite <-H3. constructor. +rewrite <-H4. rewrite <-H7. apply VertexSet.eq_refl. +auto. + +apply fold_left_assoc_map. +unfold EqualSetMap. unfold f'. unfold f. +intros. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +constructor. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). +rewrite (MapFacts.find_o _ (Vertex.eq_sym e2)). +apply VertexSet.eq_refl. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s2); constructor; apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f'. unfold f. +destruct (Vertex.eq_dec x0 a1). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H1. auto. auto. + +unfold adj_set in H. rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)) in H. +clear IHl H1 H6 H4 H5 H3. +unfold tmp in H0. +induction l. simpl in H0. +rewrite H0 in H. elim (VertexSet.empty_1 H). + +cut (EqualSetMap (fold_left f' (a0 :: l) m) (f' (fold_left f' l m) a0)). intro. +generalize (H1 y). clear H1. intro H1. simpl in H0. inversion H1. +set (tmp' := fold_left f' l m) in *. +unfold f' in H4. unfold f in H4. +destruct (Vertex.eq_dec y a0). +rewrite MapFacts.add_eq_o in H4. congruence. intuition. +rewrite MapFacts.add_neq_o in H4. +apply IHl. auto. auto. + +congruence. + +apply fold_left_assoc_map. +unfold EqualSetMap. unfold f'. unfold f. +intros. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +constructor. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). +rewrite (MapFacts.find_o _ (Vertex.eq_sym e2)). +apply VertexSet.eq_refl. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s2); constructor; apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f'. unfold f. +destruct (Vertex.eq_dec x0 a1). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H1. auto. auto. + +auto. +intuition. +rewrite MapFacts.add_neq_o in H2. +rewrite <-H2 in *. rewrite <-H1 in *. + +inversion HH0. +unfold f' in H6. unfold f in H6. +destruct (Vertex.eq_dec x a). +rewrite MapFacts.add_eq_o in H6. congruence. intuition. +rewrite MapFacts.add_neq_o in H6. +rewrite <-H6 in IHl. +apply IHl. rewrite <-H3. assumption. +auto. + +unfold f' in H5. unfold f in H5. +destruct (Vertex.eq_dec x a). +rewrite MapFacts.add_eq_o in H5. +clear H0 HH0. + +inversion H5. subst. clear H5. +rewrite H6. +apply VertexSet.add_2. +apply VertexSet.remove_2. +auto. +apply Hsym. +rewrite <-e0. +rewrite H3 in H. clear IHl H1 H4 H6. +unfold tmp in H2. +assert (eq_set_option (Some s'0) (VertexMap.find y (fold_left f' l m))). +rewrite <-H2. constructor. apply VertexSet.eq_refl. +generalize H0. clear H0 H2. intro H2. +induction l. simpl in H2. +unfold adj_set. inversion H2. subst. rewrite <-H4. assumption. + +cut (EqualSetMap (fold_left f' (a0 :: l) m) (f' (fold_left f' l m) a0)). intro. +generalize (H0 y). clear H0. simpl in H2. intro H0. inversion H0. clear H0. +rewrite <-H4 in H2. inversion H2. +set (tmp' := fold_left f' l m) in *. +unfold f' in H4. unfold f in H4. +destruct (Vertex.eq_dec y a0). +rewrite MapFacts.add_eq_o in H4. inversion H4. subst. clear H4. +inversion H2. subst. rewrite <-H6 in H1. inversion H1. subst. +apply VertexSet.remove_3 with (x:=snd_ext e). +apply VertexSet.add_3 with (x:= fst_ext e). +auto. +unfold adj_set. rewrite (MapFacts.find_o _ e1). rewrite <-H5. rewrite <-H7. assumption. +intuition. +rewrite MapFacts.add_neq_o in H4. +apply IHl. +rewrite <-H4. constructor. rewrite <-H1 in H2. inversion H2. subst. +rewrite H8. assumption. +auto. + +apply fold_left_assoc_map. +unfold EqualSetMap. unfold f'. unfold f. +intros. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +constructor. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). +rewrite (MapFacts.find_o _ (Vertex.eq_sym e2)). +apply VertexSet.eq_refl. intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s2); constructor; apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f'. unfold f. +destruct (Vertex.eq_dec x0 a1). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H0. auto. auto. +intuition. +rewrite MapFacts.add_neq_o in H5. +rewrite <-H5 in *. +rewrite H6. apply IHl. +rewrite <-H3. assumption. +auto. +auto. + +apply fold_left_assoc_map. +unfold EqualSetMap. unfold f'. unfold f. +intros. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +constructor. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). +apply VertexSet.eq_refl. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s0); constructor; apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f'. unfold f. +destruct (Vertex.eq_dec x0 a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H0. auto. auto. + +auto. +auto. +auto. +auto. + +clear H. +unfold s. unfold s'. intros z H Hne. +apply VertexSet.union_3. apply VertexSet.remove_2; auto. +Qed. + +Lemma not_eq_extremities_merge_map : forall e g, +aff_edge e -> +forall x y, +VertexSet.In x (adj_set y (imap_merge e g)) \/ +VertexSet.In x (adj_set y (pmap_merge e g (imap_merge e g))) -> +~Vertex.eq x y. + +Proof. +intros. destruct H0. +apply (not_eq_extremities_map_merge x y e (imap g)). +intros. apply (not_eq_extremities g). left. assumption. assumption. +apply (not_eq_extremities_map_merge x y e (pmap g)). +intros. apply (not_eq_extremities g). right. assumption. +apply pmap_merge_sub. assumption. +Qed. + +Lemma sym_merge : forall e g, +aff_edge e -> +In_graph_edge e g -> +forall x y, +VertexSet.In x (adj_set y (map_merge e (pmap g))) -> +VertexSet.In y (adj_set x (map_merge e (pmap g))). + +Proof. +intros e g Haff Hin x y H. +apply sym_map_merge_map with (g:=g); auto. +apply (sym_pmap g). +intros. apply (not_eq_extremities_map_merge _ _ e (pmap g)). +intros. apply (not_eq_extremities g). right. auto. +assumption. + +assert (~Vertex.eq y (snd_ext e)) as Hy. +intro. +generalize (extremities_in_merge_map e g Hin y). +rewrite MapFacts.in_find_iff. intro. +unfold adj_set in H. +case_eq (VertexMap.find y (map_merge e (pmap g))); intros. +rewrite H2 in H1. +assert (VertexSet.In y (VertexSet.remove (snd_ext e) (V g))). +rewrite <-H1. congruence. +elim (VertexSet.remove_1 (Vertex.eq_sym H0) H3). +rewrite H2 in H. elim (VertexSet.empty_1 H). + +intro. +unfold map_merge in H. + +set (f := (fun (y : VertexSet.elt) (m' : VertexMap.t VertexSet.t) => + VertexMap.add y + (VertexSet.add (fst_ext e) + (VertexSet.remove (snd_ext e) (adj_set y (pmap g)))) + m')) in *. +set (s := (VertexSet.union + (VertexSet.remove (snd_ext e) + (adj_set (fst_ext e) (pmap g))) + (VertexSet.remove (fst_ext e) + (adj_set (snd_ext e) (pmap g))))) in *. +set (s' := adj_set (snd_ext e) (pmap g)) in *. + +unfold adj_set in H. +rewrite MapFacts.remove_neq_o in H. +destruct (Vertex.eq_dec y (fst_ext e)). +rewrite MapFacts.add_eq_o in H. +unfold s in H. +destruct (VertexSet.union_1 H). +elim (VertexSet.remove_1 (Vertex.eq_sym H0) H1). +generalize (VertexSet.remove_3 H1). clear H1. intro H1. +unfold s' in H1. + +elim (not_eq_extremities g x (snd_ext e)). +right. assumption. +assumption. +intuition. +rewrite MapFacts.add_neq_o in H. + +rewrite VertexSet.fold_1 in H. +set (f' := fun a e => f e a) in *. +generalize VertexSet.elements_1. intro. +generalize (H1 (VertexSet.remove (fst_ext e) s') y). clear H1. intro H1. +induction (VertexSet.elements (VertexSet.remove (fst_ext e) s')). +simpl in H. +unfold s' in H1. +assert (InA Vertex.eq y nil). +apply H1. +apply VertexSet.remove_2. auto. +apply (sym_pmap g). rewrite <-H0. assumption. inversion H2. + +cut (EqualSetMap (fold_left f' (a :: l) (pmap g)) (f' (fold_left f' l (pmap g)) a)). intro. +generalize (H2 y). clear H2. intro H2. simpl in H. inversion H2; clear H2. +rewrite <-H4 in *. +elim (VertexSet.empty_1 H). +set (tmp := fold_left f' l (pmap g)) in *. +unfold f' in H4. unfold f in H4. +destruct (Vertex.eq_dec y a). +rewrite MapFacts.add_eq_o in H4. +rewrite <-H3 in *. clear H3. inversion H4. subst. clear H4. +rewrite H5 in H. clear H5. +destruct (proj1 (Props.Dec.F.add_iff _ _ _) H). +elim (In_graph_edge_diff_ext _ _ Hin). +apply Vertex.eq_trans with (y := x); auto. +elim (VertexSet.remove_1 (Vertex.eq_sym H0) H2). +intuition. +rewrite MapFacts.add_neq_o in H4. rewrite <-H4 in IHl. +apply IHl. +rewrite <-H5. rewrite <-H3 in H. assumption. +intro. generalize (H1 H2). clear H2. intro H2. +inversion H2; subst. +elim (n0 H7). +assumption. +auto. + +apply fold_left_assoc_map. +unfold EqualSetMap. unfold f'. unfold f. +intros. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +constructor. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). +apply VertexSet.eq_refl. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s0); constructor; apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f'. unfold f. +destruct (Vertex.eq_dec x0 a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H2. auto. auto. + +auto. +auto. + +intro. +generalize (extremities_in_merge_map e g Hin y). +rewrite MapFacts.in_find_iff. intro. +unfold adj_set in H. +case_eq (VertexMap.find y (map_merge e (pmap g))); intros. +rewrite H2 in H1. +assert (VertexSet.In y (VertexSet.remove (snd_ext e) (V g))). +rewrite <-H1. congruence. +elim (VertexSet.remove_1 (Vertex.eq_sym H0) H3). +rewrite H2 in H. elim (VertexSet.empty_1 H). +Qed. + +Lemma sym_imap_merge_map : forall e g, +aff_edge e -> +In_graph_edge e g -> +forall x y, +VertexSet.In x (adj_set y (imap_merge e g)) -> +VertexSet.In y (adj_set x (imap_merge e g)). + +Proof. +intros e g Haff Hin x y H. +apply sym_map_merge_map with (g:=g); auto. +apply (sym_imap g). +intros. apply (not_eq_extremities_merge_map e g Haff). left. assumption. + +assert (~Vertex.eq y (snd_ext e)) as Hy. +intro. +generalize (extremities_in_merge_imap e g Hin y). +rewrite MapFacts.in_find_iff. intro. +unfold adj_set in H. +case_eq (VertexMap.find y (imap_merge e g)); intros. +rewrite H2 in H1. +assert (VertexSet.In y (VertexSet.remove (snd_ext e) (V g))). +rewrite <-H1. congruence. +elim (VertexSet.remove_1 (Vertex.eq_sym H0) H3). +rewrite H2 in H. elim (VertexSet.empty_1 H). + +intro. +unfold imap_merge in H. +unfold map_merge in H. + +set (f := (fun (y : VertexSet.elt) (m' : VertexMap.t VertexSet.t) => + VertexMap.add y + (VertexSet.add (fst_ext e) + (VertexSet.remove (snd_ext e) (adj_set y (imap g)))) + m')) in *. +set (s := (VertexSet.union + (VertexSet.remove (snd_ext e) + (adj_set (fst_ext e) (imap g))) + (VertexSet.remove (fst_ext e) + (adj_set (snd_ext e) (imap g))))) in *. +set (s' := adj_set (snd_ext e) (imap g)) in *. + +unfold adj_set in H. +rewrite MapFacts.remove_neq_o in H. +destruct (Vertex.eq_dec y (fst_ext e)). +rewrite MapFacts.add_eq_o in H. +unfold s in H. +destruct (VertexSet.union_1 H). +elim (VertexSet.remove_1 (Vertex.eq_sym H0) H1). +generalize (VertexSet.remove_3 H1). clear H1. intro H1. +unfold s' in H1. + +elim (not_eq_extremities g x (snd_ext e)). +left. assumption. +assumption. +intuition. +rewrite MapFacts.add_neq_o in H. + +rewrite VertexSet.fold_1 in H. +set (f' := fun a e => f e a) in *. +generalize VertexSet.elements_1. intro. +generalize (H1 (VertexSet.remove (fst_ext e) s') y). clear H1. intro H1. +induction (VertexSet.elements (VertexSet.remove (fst_ext e) s')). +simpl in H. +unfold s' in H1. +assert (InA Vertex.eq y nil). +apply H1. +apply VertexSet.remove_2. auto. +apply (sym_imap g). rewrite <-H0. assumption. inversion H2. + +cut (EqualSetMap (fold_left f' (a :: l) (imap g)) (f' (fold_left f' l (imap g)) a)). intro. +generalize (H2 y). clear H2. intro H2. simpl in H. inversion H2; clear H2. +rewrite <-H4 in *. +elim (VertexSet.empty_1 H). +set (tmp := fold_left f' l (imap g)) in *. +unfold f' in H4. unfold f in H4. +destruct (Vertex.eq_dec y a). +rewrite MapFacts.add_eq_o in H4. +rewrite <-H3 in *. clear H3. inversion H4. subst. clear H4. +rewrite H5 in H. clear H5. +destruct (proj1 (Props.Dec.F.add_iff _ _ _) H). +elim (In_graph_edge_diff_ext _ _ Hin). +apply Vertex.eq_trans with (y := x); auto. +elim (VertexSet.remove_1 (Vertex.eq_sym H0) H2). +intuition. +rewrite MapFacts.add_neq_o in H4. rewrite <-H4 in IHl. +apply IHl. +rewrite <-H5. rewrite <-H3 in H. assumption. +intro. generalize (H1 H2). clear H2. intro H2. +inversion H2; subst. +elim (n0 H7). +assumption. +auto. + +apply fold_left_assoc_map. +unfold EqualSetMap. unfold f'. unfold f. +intros. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +constructor. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). +apply VertexSet.eq_refl. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s0); constructor; apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f'. unfold f. +destruct (Vertex.eq_dec x0 a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H2. auto. auto. + +auto. +auto. + +intro. +generalize (extremities_in_merge_imap e g Hin y). +rewrite MapFacts.in_find_iff. intro. +unfold adj_set in H. +case_eq (VertexMap.find y (imap_merge e g)); intros. +rewrite H2 in H1. +assert (VertexSet.In y (VertexSet.remove (snd_ext e) (V g))). +rewrite <-H1. congruence. +elim (VertexSet.remove_1 (Vertex.eq_sym H0) H3). +rewrite H2 in H. elim (VertexSet.empty_1 H). +Qed. + +Lemma pmap_merge_in_merge : forall x e g, +aff_edge e -> +In_graph_edge e g -> +VertexMap.In x (pmap_merge e g (imap_merge e g)) -> +VertexMap.In x (map_merge e (pmap g)). + +Proof. +intros x e g Haff HH H. +unfold pmap_merge in H. +unfold resolve_conflicts in H. +set (f := (fun (x : VertexSet.elt) (m : VertexMap.t VertexSet.t) => + VertexMap.add x + (VertexSet.remove (fst_ext e) + (adj_set x (map_merge e (pmap g)))) m)) in *. +rewrite MapFacts.in_find_iff in *. intro H0. +destruct (Vertex.eq_dec x (fst_ext e)). +rewrite (MapFacts.find_o _ e0) in H0. +assert (VertexMap.In (fst_ext e) (map_merge e (pmap g))). +apply pmap_merge_domain_2. assumption. +apply VertexSet.remove_2. apply (In_graph_edge_diff_ext _ _ HH). +apply (proj1 (In_graph_edge_in_ext _ _ HH)). +rewrite MapFacts.in_find_iff in H1. congruence. +rewrite MapFacts.add_neq_o in H. +set (s' := (VertexSet.inter (adj_set (fst_ext e) (map_merge e (pmap g))) + (adj_set (fst_ext e) (imap_merge e g)))) in *. +set (m := map_merge e (pmap g)) in *. +rewrite VertexSet.fold_1 in H. +set (f' := fun a e => f e a) in *. +assert (forall z, InA Vertex.eq z (VertexSet.elements s') -> + VertexMap.In z m) as Hin. +intros. +apply pmap_merge_domain_2. assumption. +generalize (VertexSet.elements_2 H1). clear H1. intro H1. +unfold s' in H1. generalize (VertexSet.inter_1 H1). +generalize (VertexSet.inter_2 H1). clear H1. intros. +rewrite <-(extremities_in_merge_imap e g HH). +generalize (sym_imap_merge_map e g Haff HH _ _ H1). intro. +rewrite MapFacts.in_find_iff. intro. +unfold adj_set in H3. rewrite H4 in H3. elim (VertexSet.empty_1 H3). +induction (VertexSet.elements s'). simpl in H. congruence. +cut (EqualSetMap (fold_left f' (a :: l) m) (f' (fold_left f' l m) a)). intro. +generalize (H1 x). clear H1. simpl in H. intro H1. inversion H1; clear H1. +congruence. +set (tmp := fold_left f' l m) in *. +unfold f' in H3. unfold f in H3. +destruct (Vertex.eq_dec x a). +rewrite MapFacts.add_eq_o in H3. inversion H3; subst; clear H3. +rewrite <-H2 in H. +rewrite (MapFacts.find_o _ e0) in H0. +assert (VertexMap.In a m). +apply Hin. left. auto. +rewrite MapFacts.in_find_iff in H1. congruence. +intuition. +rewrite MapFacts.add_neq_o in H3. +apply IHl. congruence. +intros. apply Hin. right. auto. +auto. + + +apply fold_left_assoc_map. +unfold EqualSetMap. unfold f'. unfold f. +intros. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y). +rewrite MapFacts.add_eq_o. +constructor. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). +apply VertexSet.eq_refl. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s); constructor; apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f'. unfold f. +destruct (Vertex.eq_dec x0 a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H1. auto. auto. +auto. +Qed. + +Lemma pmap_merge_merge_in : forall x e g, +In_graph_edge e g -> +VertexMap.In x (map_merge e (pmap g)) -> +VertexMap.In x (pmap_merge e g (imap_merge e g)). + +Proof. +intros x e g HH H. +unfold pmap_merge. +unfold resolve_conflicts. +set (f := (fun (x0 : VertexSet.elt) (m : VertexMap.t VertexSet.t) => + VertexMap.add x0 + (VertexSet.remove (fst_ext e) (adj_set x0 (map_merge e (pmap g)))) + m)) in *. +set (s := (VertexSet.inter (adj_set (fst_ext e) (map_merge e (pmap g))) + (adj_set (fst_ext e) (imap_merge e g)))) in *. +rewrite MapFacts.in_find_iff. +destruct (Vertex.eq_dec x (fst_ext e)). +rewrite MapFacts.add_eq_o. congruence. +intuition. +rewrite MapFacts.add_neq_o. +intro H0. +rewrite VertexSet.fold_1 in H0. +set (f' := fun a e => f e a) in *. +assert (VertexMap.find x (map_merge e (pmap g)) = None). +induction (VertexSet.elements s). simpl in H0. assumption. +set (m := map_merge e (pmap g)) in *. +cut (EqualSetMap (fold_left f' (a :: l) m) (f' (fold_left f' l m) a)). intro. +generalize (H1 x). clear H1. intro H1. simpl in H0. inversion H1; clear H1. +set (tmp := fold_left f' l m) in *. +unfold f' in H4. unfold f in H4. +destruct (Vertex.eq_dec x a). rewrite MapFacts.add_eq_o in H4. congruence. +intuition. +rewrite MapFacts.add_neq_o in H4. +apply IHl. auto. +auto. + +congruence. +apply fold_left_assoc_map. +unfold EqualSetMap. unfold f'. unfold f. +intros. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y). +rewrite MapFacts.add_eq_o. +constructor. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). +apply VertexSet.eq_refl. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s0); constructor; apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f'. unfold f. +destruct (Vertex.eq_dec x0 a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H1. auto. auto. +auto. + +rewrite MapFacts.in_find_iff in H. congruence. +auto. +Qed. + +Lemma extremities_in_merge_pmap e g : +aff_edge e -> +In_graph_edge e g -> +(forall x, +VertexMap.In x (pmap_merge e g (imap_merge e g)) <-> +VertexSet.In x (VertexSet.remove (snd_ext e) (V g))). + +Proof. +intros e g Haff HH x;split; intro H0. +apply pmap_merge_domain_1. assumption. +apply pmap_merge_in_merge. assumption. +assumption. assumption. +apply pmap_merge_merge_in. assumption. +apply pmap_merge_domain_2. assumption. +assumption. +Qed. + +Lemma merge_conflicts_aux_1 : forall e g, +aff_edge e -> +In_graph_edge e g -> +forall x y, +VertexSet.In x (adj_set y (imap_merge e g)) -> +~Vertex.eq y (fst_ext e) -> +~Vertex.eq x (fst_ext e) -> +VertexSet.In x (adj_set y (imap g)). + +Proof. +intros. +unfold imap_merge in H1. +unfold map_merge in H1. +set (m := imap g) in *. +set (f := fun (y : VertexSet.elt) (m' : VertexMap.t VertexSet.t) => + VertexMap.add y + (VertexSet.add (fst_ext e) + (VertexSet.remove (snd_ext e) (adj_set y m))) + m') in *. +set (s := (VertexSet.union + (VertexSet.remove (snd_ext e) + (adj_set (fst_ext e) m)) + (VertexSet.remove (fst_ext e) + (adj_set (snd_ext e) m)))) in *. +set (s' := (adj_set (snd_ext e) m)) in *. +unfold adj_set in H1. +destruct (Vertex.eq_dec y (snd_ext e)). +rewrite MapFacts.remove_eq_o in H1. elim (VertexSet.empty_1 H1). +intuition. +rewrite MapFacts.remove_neq_o in H1. rewrite MapFacts.add_neq_o in H1. +rewrite VertexSet.fold_1 in H1. set (f' := fun a e => f e a) in *. +induction (VertexSet.elements (VertexSet.remove (fst_ext e) s')). simpl in H1. +assumption. +cut (EqualSetMap (fold_left f' (a :: l) m) (f' (fold_left f' l m) a)). intro. +generalize (H4 y). clear H4. intro H4. simpl in H1. inversion H4; subst; clear H4. +rewrite <-H6 in H1. elim (VertexSet.empty_1 H1). +set (tmp := fold_left f' l m) in *. +unfold f' in H6. unfold f in H6. +destruct (Vertex.eq_dec y a). rewrite MapFacts.add_eq_o in H6. +rewrite <-H5 in H1. rewrite H7 in H1. +inversion H6; subst; clear H6. +destruct (proj1 (Props.Dec.F.add_iff _ _ _) H1). elim H3; auto. +generalize (VertexSet.remove_3 H4). intro. +unfold adj_set. rewrite (MapFacts.find_o _ e0). assumption. +intuition. +rewrite MapFacts.add_neq_o in H6. +rewrite <-H6 in IHl. +apply IHl. rewrite <-H7. rewrite <-H5 in H1. assumption. +auto. + +apply fold_left_assoc_map. +unfold EqualSetMap. unfold f'. unfold f. +intros. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +constructor. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). +apply VertexSet.eq_refl. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s0); constructor; apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f'. unfold f. +destruct (Vertex.eq_dec x0 a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H4. auto. auto. + +auto. +auto. +Qed. + +Lemma merge_conflicts_aux_2 : forall e g, +aff_edge e -> +In_graph_edge e g -> +forall x y, +VertexSet.In x (adj_set y (map_merge e (pmap g))) -> +~Vertex.eq y (fst_ext e) -> +~Vertex.eq x (fst_ext e) -> +VertexSet.In x (adj_set y (pmap g)). + +Proof. +intros. +unfold map_merge in H1. +set (m := pmap g) in *. +set (f := fun (y : VertexSet.elt) (m' : VertexMap.t VertexSet.t) => + VertexMap.add y + (VertexSet.add (fst_ext e) + (VertexSet.remove (snd_ext e) (adj_set y m))) + m') in *. +set (s := (VertexSet.union + (VertexSet.remove (snd_ext e) + (adj_set (fst_ext e) m)) + (VertexSet.remove (fst_ext e) + (adj_set (snd_ext e) m)))) in *. +set (s' := (adj_set (snd_ext e) m)) in *. +unfold adj_set in H1. +destruct (Vertex.eq_dec y (snd_ext e)). +rewrite MapFacts.remove_eq_o in H1. elim (VertexSet.empty_1 H1). +intuition. +rewrite MapFacts.remove_neq_o in H1. rewrite MapFacts.add_neq_o in H1. +rewrite VertexSet.fold_1 in H1. set (f' := fun a e => f e a) in *. +induction (VertexSet.elements (VertexSet.remove (fst_ext e) s')). simpl in H1. +assumption. +cut (EqualSetMap (fold_left f' (a :: l) m) (f' (fold_left f' l m) a)). intro. +generalize (H4 y). clear H4. intro H4. simpl in H1. inversion H4; subst; clear H4. +rewrite <-H6 in H1. elim (VertexSet.empty_1 H1). +set (tmp := fold_left f' l m) in *. +unfold f' in H6. unfold f in H6. +destruct (Vertex.eq_dec y a). rewrite MapFacts.add_eq_o in H6. +rewrite <-H5 in H1. rewrite H7 in H1. +inversion H6; subst; clear H6. +destruct (proj1 (Props.Dec.F.add_iff _ _ _) H1). elim H3; auto. +generalize (VertexSet.remove_3 H4). intro. +unfold adj_set. rewrite (MapFacts.find_o _ e0). assumption. +intuition. +rewrite MapFacts.add_neq_o in H6. +rewrite <-H6 in IHl. +apply IHl. rewrite <-H7. rewrite <-H5 in H1. assumption. +auto. + +apply fold_left_assoc_map. +unfold EqualSetMap. unfold f'. unfold f. +intros. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +constructor. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). +apply VertexSet.eq_refl. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s0); constructor; apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f'. unfold f. +destruct (Vertex.eq_dec x0 a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H4. auto. auto. + +auto. +auto. +Qed. + +Lemma merge_conflicts : forall e g, +aff_edge e -> +In_graph_edge e g -> +forall x y, +VertexSet.In x (adj_set y (imap_merge e g)) -> +VertexSet.In x (adj_set y (map_merge e (pmap g))) -> +~Vertex.eq y (fst_ext e) -> +~Vertex.eq x (fst_ext e) -> +False. + +Proof. +intros. +apply (simple_graph g x y). +split. +apply (merge_conflicts_aux_1 e g H H0 x y H1 H3 H4). +apply (merge_conflicts_aux_2 e g H H0 x y H2 H3 H4). +Qed. + +Lemma resolve_conflicts_map_3 : forall x y e g, +aff_edge e -> +In_graph_edge e g -> +VertexSet.In x (adj_set y (imap_merge e g)) /\ +VertexSet.In x (adj_set y (resolve_conflicts (fst_ext e) (map_merge e (pmap g)) + (adj_set (fst_ext e) (map_merge e (pmap g))) + (adj_set (fst_ext e) (imap_merge e g)))) + +-> False. + +Proof. +intros x y e g Haff Hin H. destruct H. +unfold resolve_conflicts in H0. +set (f := (fun (x : VertexSet.elt) (m : VertexMap.t VertexSet.t) => + VertexMap.add x + (VertexSet.remove (fst_ext e) + (adj_set x (map_merge e (pmap g)))) m)) in *. +set (s := (VertexSet.diff (adj_set (fst_ext e) (map_merge e (pmap g))) + (adj_set (fst_ext e) (imap_merge e g)))) in *. +set (s' := (VertexSet.inter (adj_set (fst_ext e) (map_merge e (pmap g))) + (adj_set (fst_ext e) (imap_merge e g)))) in *. +rewrite VertexSet.fold_1 in H0. +set (f' := fun a e => f e a) in *. +generalize VertexSet.elements_1. intro HH. +generalize (HH s' y). clear HH. intro HH. +induction (VertexSet.elements s'). simpl in H0. +destruct (Vertex.eq_dec y (fst_ext e)). +unfold adj_set in H0. rewrite MapFacts.add_eq_o in H0. +elim (VertexSet.diff_2 H0). +unfold adj_set. rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). assumption. +intuition. +unfold adj_set in H0. rewrite MapFacts.add_neq_o in H0. + +assert (InA Vertex.eq y nil). +apply HH. +unfold s'. apply VertexSet.inter_3. +cut (Vertex.eq x (fst_ext e)). intro. +unfold adj_set. rewrite (MapFacts.find_o _ (Vertex.eq_sym H1)). +apply sym_merge; auto. +destruct (Vertex.eq_dec x (fst_ext e)). +intuition. +apply False_ind. apply (merge_conflicts e g Haff Hin x y); auto. + +destruct (Vertex.eq_dec x (fst_ext e)). +unfold adj_set. rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +apply (sym_imap_merge_map e g Haff Hin). assumption. +apply False_ind. +apply (merge_conflicts e g Haff Hin x y); auto. inversion H1. +auto. + +set (m := map_merge e (pmap g)) in *. +cut (EqualSetMap (fold_left f' (a :: l) m) (f' (fold_left f' l m) a)). intro. +generalize (H1 y). clear H1. simpl in H0. intro H1. inversion H1; clear H1. +set (tmp := fold_left f' l m) in *. +unfold f' in H4. unfold f in H4. +destruct (Vertex.eq_dec y a). rewrite MapFacts.add_eq_o in H4. congruence. +intuition. +rewrite MapFacts.add_neq_o in H4. +destruct (Vertex.eq_dec y (fst_ext e)). +unfold adj_set in H0. rewrite MapFacts.add_eq_o in H0. +unfold adj_set in IHl. rewrite MapFacts.add_eq_o in IHl. +apply IHl. assumption. intros. +generalize (HH H1). intro H2. inversion H2; subst. +elim (n H6). auto. +intuition. +intuition. + +unfold adj_set in H0. rewrite MapFacts.add_neq_o in H0. +rewrite <-H3 in H0. elim (VertexSet.empty_1 H0). +auto. +auto. + +set (tmp := fold_left f' l m) in *. +unfold f' in H3. unfold f in H3. +destruct (Vertex.eq_dec y a). rewrite MapFacts.add_eq_o in H3. +destruct (Vertex.eq_dec y (fst_ext e)). +unfold adj_set in H0. +rewrite MapFacts.add_eq_o in H0. +elim (VertexSet.diff_2 H0). +unfold adj_set. rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). assumption. +intuition. +unfold adj_set in H0. rewrite MapFacts.add_neq_o in H0. +rewrite <-H2 in H0. +inversion H3; subst; clear H3. +rewrite H4 in H0. clear H4. +destruct (Vertex.eq_dec x (fst_ext e)). +elim (VertexSet.remove_1 (Vertex.eq_sym e1) H0). +apply (merge_conflicts e g Haff Hin x y); auto. +unfold adj_set. rewrite (MapFacts.find_o _ e0). apply (VertexSet.remove_3 H0). +auto. +intuition. + +rewrite MapFacts.add_neq_o in H3. +unfold adj_set in IHl. +destruct (Vertex.eq_dec y (fst_ext e)). +rewrite MapFacts.add_eq_o in IHl. +unfold adj_set in H0. +rewrite MapFacts.add_eq_o in H0. +apply IHl. assumption. intro. +generalize (HH H1). intro. +inversion H5; subst. +elim (n H7). +assumption. +intuition. +intuition. +rewrite MapFacts.add_neq_o in IHl. rewrite <-H3 in IHl. +unfold adj_set in H0. rewrite MapFacts.add_neq_o in H0. +rewrite <-H2 in H0. +apply IHl. rewrite <-H4. assumption. +intro. generalize (HH H1). intro. +inversion H5; subst. +elim (n H7). +assumption. +auto. +auto. +auto. + +apply fold_left_assoc_map. +unfold EqualSetMap. unfold f'. unfold f. +intros. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +constructor. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). +apply VertexSet.eq_refl. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s0); constructor; apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f'. unfold f. +destruct (Vertex.eq_dec x0 a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H1. auto. auto. +Qed. + +Lemma simple_graph_merge_map : forall e g, +aff_edge e -> +In_graph_edge e g -> +forall x y, +VertexSet.In x (adj_set y (imap_merge e g)) /\ +VertexSet.In x (adj_set y (pmap_merge e g (imap_merge e g))) -> +False. + +Proof. +exact (fun e g H H0 x y H1 => + (resolve_conflicts_map_3 x y e g H H0 H1)). +Qed. + +Lemma sym_map_merge_pmap : forall e g x y, +aff_edge e -> +In_graph_edge e g -> +VertexSet.In x (adj_set y (map_merge e (pmap g))) -> +VertexSet.In y (adj_set x (map_merge e (pmap g))). + +Proof. +intros e g x y Haff Hin H. +apply sym_map_merge_map with (g:=g); auto. +apply (sym_pmap g). +intros. +apply (not_eq_extremities_map_merge a b e (pmap g)). +intros. apply (not_eq_extremities g). right. assumption. assumption. + +assert (~Vertex.eq y (snd_ext e)) as Hy. +intro. +assert (VertexSet.In y (VertexSet.remove (snd_ext e) (V g))). +apply (pmap_merge_domain_1 y _ _ Hin). +rewrite MapFacts.in_find_iff. intro. +unfold adj_set in H. rewrite H1 in H. elim (VertexSet.empty_1 H). +elim (VertexSet.remove_1 (Vertex.eq_sym H0) H1). + +intro. +unfold map_merge in H. + +set (f := (fun (y : VertexSet.elt) (m' : VertexMap.t VertexSet.t) => + VertexMap.add y + (VertexSet.add (fst_ext e) + (VertexSet.remove (snd_ext e) (adj_set y (pmap g)))) + m')) in *. +set (s := (VertexSet.union + (VertexSet.remove (snd_ext e) + (adj_set (fst_ext e) (pmap g))) + (VertexSet.remove (fst_ext e) + (adj_set (snd_ext e) (pmap g))))) in *. +set (s' := adj_set (snd_ext e) (pmap g)) in *. + +unfold adj_set in H. +rewrite MapFacts.remove_neq_o in H. +destruct (Vertex.eq_dec y (fst_ext e)). +rewrite MapFacts.add_eq_o in H. +unfold s in H. +destruct (VertexSet.union_1 H). +elim (VertexSet.remove_1 (Vertex.eq_sym H0) H1). +generalize (VertexSet.remove_3 H1). clear H1. intro H1. +unfold s' in H1. + +elim (not_eq_extremities g x (snd_ext e)). +right. assumption. +assumption. +intuition. +rewrite MapFacts.add_neq_o in H. + +rewrite VertexSet.fold_1 in H. +set (f' := fun a e => f e a) in *. +generalize VertexSet.elements_1. intro. +generalize (H1 (VertexSet.remove (fst_ext e) s') y). clear H1. intro H1. +induction (VertexSet.elements (VertexSet.remove (fst_ext e) s')). +simpl in H. +unfold s' in H1. +assert (InA Vertex.eq y nil). +apply H1. +apply VertexSet.remove_2. auto. +apply (sym_pmap g). rewrite <-H0. assumption. inversion H2. + +cut (EqualSetMap (fold_left f' (a :: l) (pmap g)) (f' (fold_left f' l (pmap g)) a)). intro. +generalize (H2 y). clear H2. intro H2. simpl in H. inversion H2; clear H2. +rewrite <-H4 in *. +elim (VertexSet.empty_1 H). +set (tmp := fold_left f' l (pmap g)) in *. +unfold f' in H4. unfold f in H4. +destruct (Vertex.eq_dec y a). +rewrite MapFacts.add_eq_o in H4. +rewrite <-H3 in *. clear H3. inversion H4. subst. clear H4. +rewrite H5 in H. clear H5. +destruct (proj1 (Props.Dec.F.add_iff _ _ _) H). +elim (In_graph_edge_diff_ext _ _ Hin). +apply Vertex.eq_trans with (y := x); auto. +elim (VertexSet.remove_1 (Vertex.eq_sym H0) H2). +intuition. +rewrite MapFacts.add_neq_o in H4. rewrite <-H4 in IHl. +apply IHl. +rewrite <-H5. rewrite <-H3 in H. assumption. +intro. generalize (H1 H2). clear H2. intro H2. +inversion H2; subst. +elim (n0 H7). +assumption. +auto. + +apply fold_left_assoc_map. +unfold EqualSetMap. unfold f'. unfold f. +intros. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +constructor. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). +apply VertexSet.eq_refl. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s0); constructor; apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f'. unfold f. +destruct (Vertex.eq_dec x0 a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H2. auto. auto. + +auto. +auto. + +intro. +assert (VertexSet.In y (VertexSet.remove (snd_ext e) (V g))). +apply (pmap_merge_domain_1 y _ _ Hin). +rewrite MapFacts.in_find_iff. intro. +unfold adj_set in H. rewrite H1 in H. elim (VertexSet.empty_1 H). +elim (VertexSet.remove_1 (Vertex.eq_sym H0) H1). +Qed. + +Lemma sym_resolve : forall x y e g, +aff_edge e -> +In_graph_edge e g -> +VertexSet.In x (adj_set y (resolve_conflicts (fst_ext e) (map_merge e (pmap g)) + (adj_set (fst_ext e) (map_merge e (pmap g))) + (adj_set (fst_ext e) (imap_merge e g)))) -> +VertexSet.In y (adj_set x (resolve_conflicts (fst_ext e) (map_merge e (pmap g)) + (adj_set (fst_ext e) (map_merge e (pmap g))) + (adj_set (fst_ext e) (imap_merge e g)))). + +Proof. +unfold resolve_conflicts; intros x y e g p q H. +set (f := (fun (x0 : VertexSet.elt) (m : VertexMap.t VertexSet.t) => + VertexMap.add x0 + (VertexSet.remove (fst_ext e) + (adj_set x0 (map_merge e (pmap g)))) m)) in *. +set (s := (VertexSet.diff (adj_set (fst_ext e) (map_merge e (pmap g))) + (adj_set (fst_ext e) (imap_merge e g)))) in *. +set (s' := (VertexSet.inter (adj_set (fst_ext e) (map_merge e (pmap g))) + (adj_set (fst_ext e) (imap_merge e g)))) in *. +set (m := map_merge e (pmap g)) in *. +rewrite VertexSet.fold_1 in *. +set (f' := fun a e => f e a) in *. +generalize VertexSet.elements_1. intro. +generalize (H0 s' y). clear H0. intro HH. +generalize VertexSet.elements_2. intro. +generalize (H0 s'). clear H0. intro HHH. +induction (VertexSet.elements s'). simpl in *. +destruct (Vertex.eq_dec x (fst_ext e)). +unfold adj_set. rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec y (fst_ext e)). +unfold adj_set in H. rewrite MapFacts.add_eq_o in H. +generalize (VertexSet.diff_1 H). intro H0. +rewrite e0 in H0. +unfold m in H0. unfold map_merge in H0. +unfold adj_set in H0. +rewrite MapFacts.remove_neq_o in H0. +rewrite MapFacts.add_eq_o in H0. +fold (adj_set (fst_ext e) (pmap g)) in H0. +fold (adj_set (snd_ext e) (pmap g)) in H0. +destruct (VertexSet.union_1 H0). +generalize (VertexSet.remove_3 H1). intro. +elim (not_eq_extremities g (fst_ext e) (fst_ext e)). right. auto. +intuition. +elim (VertexSet.remove_1 (Vertex.eq_refl _) H1). +intuition. +apply (In_graph_edge_diff_ext _ _ q). intuition. +unfold adj_set in H. rewrite MapFacts.add_neq_o in H. +fold (adj_set y m) in H. +destruct (Props.In_dec y (adj_set (fst_ext e) (imap_merge e g))). +assert (InA Vertex.eq y nil). +apply HH. apply VertexSet.inter_3. +unfold adj_set. rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +apply (sym_map_merge_pmap e g _ _ p q H). assumption. inversion H0. +apply VertexSet.diff_3. +unfold adj_set. rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +apply (sym_map_merge_pmap e g _ _ p q H). assumption. +auto. +intuition. + +unfold adj_set. rewrite MapFacts.add_neq_o. +fold (adj_set x m). +destruct (Vertex.eq_dec y (fst_ext e)). +unfold adj_set in H. rewrite MapFacts.add_eq_o in H. +apply (sym_map_merge_pmap e g _ _ p q). +unfold adj_set. rewrite (MapFacts.find_o _ e0). +apply (VertexSet.diff_1 H). +intuition. +unfold adj_set in H. rewrite MapFacts.add_neq_o in H. +apply (sym_map_merge_pmap e g _ _ p q H). +auto. +auto. + +cut (EqualSetMap (fold_left f' (a :: l) m) (f' (fold_left f' l m) a)). intro. +destruct (Vertex.eq_dec x (fst_ext e)). +unfold adj_set. rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec y (fst_ext e)). +unfold adj_set in H. rewrite MapFacts.add_eq_o in H. +rewrite e1. rewrite <-e0. assumption. +intuition. +unfold adj_set in H. rewrite MapFacts.add_neq_o in H. +generalize (H0 y). simpl in H. intro H1. inversion H1; clear H1. +rewrite <-H3 in H. +elim (VertexSet.empty_1 H). +rewrite <-H2 in H. clear H2. +set (tmp := fold_left f' l m) in *. +unfold f' in H3. unfold f in H3. +destruct (Vertex.eq_dec y a). rewrite MapFacts.add_eq_o in H3. +inversion H3; subst; clear H3. +rewrite H4 in H. elim (VertexSet.remove_1 (Vertex.eq_sym e0) H). +intuition. +rewrite MapFacts.add_neq_o in H3. +unfold adj_set in IHl. rewrite MapFacts.add_neq_o in IHl. + rewrite MapFacts.add_eq_o in IHl. +rewrite <-H3 in IHl. +apply IHl. rewrite <-H4. assumption. +intros. generalize (HH H1). intro. +inversion H2; subst. +elim (n0 H6). +assumption. +auto. +intuition. +auto. +auto. +auto. +intuition. + +unfold adj_set. rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec y (fst_ext e)). +unfold adj_set in H. rewrite MapFacts.add_eq_o in H. +generalize (H0 x). intro H1. inversion H1; subst; clear H1. +set (tmp := fold_left f' l m) in *. +unfold f' in H4. unfold f in H4. +destruct (Vertex.eq_dec x a). +rewrite MapFacts.add_eq_o in H4. congruence. intuition. +destruct (Vertex.eq_dec y a). +assert (VertexSet.In y s'). apply HHH. +left. auto. +rewrite e0 in H1. +unfold s' in H1. +elim (not_eq_extremities_merge_map e g p (fst_ext e) (fst_ext e)). +left. apply (VertexSet.inter_2 H1). +intuition. +rewrite MapFacts.add_neq_o in H4. +unfold adj_set in IHl. rewrite MapFacts.add_eq_o in IHl. + rewrite MapFacts.add_neq_o in IHl. +rewrite <-H4 in IHl. +assert (VertexSet.In y VertexSet.empty). +apply IHl. assumption. +intro. generalize (HH H1). intro. +inversion H2; subst. +elim (n1 H6). auto. +intros. apply HHH. right. auto. +elim (VertexSet.empty_1 H1). +auto. +intuition. +auto. + +unfold adj_set in IHl. rewrite MapFacts.add_eq_o in IHl. + rewrite MapFacts.add_neq_o in IHl. +simpl. rewrite <-H2. rewrite H4. +set (tmp := fold_left f' l m) in *. +unfold f' in H3. unfold f in H3. +destruct (Vertex.eq_dec x a). +rewrite MapFacts.add_eq_o in H3. +inversion H3; subst; clear H3. +assert (VertexSet.In x s'). +apply HHH. left. auto. +elim (VertexSet.diff_2 H (VertexSet.inter_2 H1)). intuition. +rewrite MapFacts.add_neq_o in H3. +rewrite <-H3 in IHl. +destruct (Vertex.eq_dec y a). +assert (VertexSet.In y s'). apply HHH. left. intuition. +unfold s' in H1. +elim (not_eq_extremities_merge_map e g p (fst_ext e) (fst_ext e)). +left. rewrite e0 in H1. apply (VertexSet.inter_2 H1). intuition. + +apply IHl. +assumption. +intro. generalize (HH H1). intro. +inversion H5; subst. +elim (n1 H7). +auto. + +intros. apply HHH. right. auto. +auto. +auto. +intuition. +intuition. + +unfold adj_set in H. rewrite MapFacts.add_neq_o in H. +simpl in H. +unfold adj_set in IHl. rewrite MapFacts.add_neq_o in IHl. + rewrite MapFacts.add_neq_o in IHl. +generalize (H0 y). intro H1. inversion H1; clear H1. +rewrite <-H3 in H. elim (VertexSet.empty_1 H). +rewrite <-H2 in H. +set (tmp := fold_left f' l m) in *. +unfold f' in H3. unfold f in H3. +destruct (Vertex.eq_dec y a). rewrite MapFacts.add_eq_o in H3. +inversion H3; subst; clear H3. +generalize H. clear H. intro H. rewrite H4 in H. +generalize (VertexSet.remove_3 H). clear H. intro H. +unfold adj_set in H. rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)) in H. +fold (adj_set y m) in H. +generalize (sym_map_merge_pmap e g _ _ p q H). clear H. intro H. +fold m in H. +clear IHl H0 H2 HH HHH. + +set (l' := a :: l) in *. +induction l'. simpl. assumption. + +(* + unfold f'. unfold f. +destruct (Vertex.eq_dec x a). rewrite MapFacts.add_eq_o. +apply VertexSet.remove_2; auto. +unfold adj_set. rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). auto. +auto. +rewrite MapFacts.add_neq_o. assumption. +auto. +*) + +cut (EqualSetMap (fold_left f' (a0 :: l') m) (f' (fold_left f' l' m) a0)). intro. +generalize (H0 x). clear H0. intro H0. simpl. inversion H0; clear H0. +set (tmp' := fold_left f' l' m) in *. +unfold f' in H3. unfold f in H3. +destruct (Vertex.eq_dec x a0). +rewrite MapFacts.add_eq_o in H3. congruence. intuition. +rewrite MapFacts.add_neq_o in H3. +rewrite <-H3 in IHl'. elim (VertexSet.empty_1 IHl'). +auto. +set (tmp' := fold_left f' l' m) in *. +unfold f' in H2. unfold f in H2. +destruct (Vertex.eq_dec x a0). +rewrite MapFacts.add_eq_o in H2. inversion H2; subst; clear H2. +rewrite H3. apply VertexSet.remove_2. auto. +unfold adj_set. rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). assumption. +intuition. +rewrite MapFacts.add_neq_o in H2. +rewrite <-H2 in IHl'. +rewrite H3. auto. +auto. + +apply fold_left_assoc_map. +unfold EqualSetMap. unfold f'. unfold f. +intros. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +constructor. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e2)). +rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). +apply VertexSet.eq_refl. intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s1); constructor; apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f'. unfold f. +destruct (Vertex.eq_dec x0 a1). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H0. auto. intuition. + +intuition. +rewrite MapFacts.add_neq_o in H3. +rewrite <-H3 in IHl. + +generalize (H0 x). intro. inversion H1; clear H1. +unfold f' in H7. unfold f in H7. +destruct (Vertex.eq_dec x a). +rewrite MapFacts.add_eq_o in H7. congruence. intuition. +rewrite MapFacts.add_neq_o in H7. +rewrite <-H7 in IHl. simpl. rewrite <-H6. +apply IHl. +rewrite <-H4. auto. + +intro. generalize (HH H1). intro. +inversion H5; subst. +elim (n1 H9). +auto. + +intros. apply HHH. right. auto. +auto. + +unfold f' in H6. unfold f in H6. +destruct (Vertex.eq_dec x a). +rewrite MapFacts.add_eq_o in H6. +inversion H6; subst; clear H6. +simpl. rewrite <-H5. +rewrite H7. +apply VertexSet.remove_2. auto. + +case_eq (VertexMap.find x tmp); intros. +rewrite H1 in IHl. +assert (VertexSet.In y t0). +apply IHl. rewrite <-H4. assumption. + +intro. generalize (HH H6). intro. +inversion H8; subst. +elim (n1 H10). +auto. + +intros. apply HHH. right. auto. +clear H HH HHH IHl H0 H4 H2 H3 H5 H7. +assert (eq_set_option (Some t0) (VertexMap.find x tmp)). +rewrite H1. constructor. apply VertexSet.eq_refl. clear H1. +unfold tmp in H. +unfold adj_set. rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +fold (adj_set x m). +induction l. simpl in H. inversion H. +unfold adj_set. rewrite <-H1. rewrite <-H2. auto. + +cut (EqualSetMap (fold_left f' (a0 :: l) m) (f' (fold_left f' l m) a0)). intro. +generalize (H0 x). clear H0. intro H0. simpl in H. inversion H0; subst. +rewrite <-H2 in H. inversion H. +set (tmp' := fold_left f' l m) in *. +unfold f' in H2. unfold f in H2. +destruct (Vertex.eq_dec x a0). +rewrite MapFacts.add_eq_o in H2. +inversion H; subst. rewrite <-H1 in H5. +inversion H5; subst; clear H5. +rewrite H7 in H6. rewrite H3 in H6. inversion H2; subst; clear H2 H3 H7. +unfold adj_set. rewrite (MapFacts.find_o _ e1). apply (VertexSet.remove_3 H6). intuition. +rewrite MapFacts.add_neq_o in H2. rewrite <-H2 in IHl. +apply IHl. constructor. rewrite <-H1 in H. inversion H; subst. +rewrite H7. auto. +auto. + +apply fold_left_assoc_map. +unfold EqualSetMap. unfold f'. unfold f. +intros. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +constructor. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e2)). +rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). +apply VertexSet.eq_refl. intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s2); constructor; apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f'. unfold f. +destruct (Vertex.eq_dec x0 a1). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H0. auto. auto. + +rewrite H1 in IHl. +assert (VertexSet.In y VertexSet.empty). +apply IHl. +rewrite <-H4. auto. +intro. generalize (HH H6). intro. +inversion H8; subst. +elim (n1 H10). +auto. + +intros. apply HHH. right. auto. +elim (VertexSet.empty_1 H6). intuition. + +rewrite MapFacts.add_neq_o in H6. +rewrite <-H6 in IHl. simpl. rewrite <-H5. +rewrite H7. apply IHl. +rewrite <-H4. auto. +intro. generalize (HH H1). intro. +inversion H8; subst. +elim (n1 H10). +auto. + +intros. apply HHH. right. auto. +auto. +auto. +auto. +auto. +auto. +auto. + +apply fold_left_assoc_map. +unfold EqualSetMap. unfold f'. unfold f. +intros. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +constructor. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). +apply VertexSet.eq_refl. intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s0); constructor; apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f'. unfold f. +destruct (Vertex.eq_dec x0 a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H0. auto. auto. +Qed. + +Lemma sym_pmap_merge_map e g : +aff_edge e -> +In_graph_edge e g -> +forall x y, +VertexSet.In x (adj_set y (pmap_merge e g (imap_merge e g))) -> +VertexSet.In y (adj_set x (pmap_merge e g (imap_merge e g))). + +Proof. +intros. apply sym_resolve; assumption. +Qed. + +Definition merge e g + (q : In_graph_edge e g) + (p : aff_edge e) := + let im := imap_merge e g in + let pm := pmap_merge e g im in + Make_Graph (VertexSet.remove (snd_ext e) (V g)) + im + pm + (extremities_in_merge_imap e g q) + (extremities_in_merge_pmap e g p q) + (simple_graph_merge_map e g p q) + (sym_imap_merge_map e g p q) + (sym_pmap_merge_map e g p q) + (not_eq_extremities_merge_map e g p). + +Lemma In_graph_dec : forall v g, +{In_graph v g} + {~In_graph v g}. + +Proof. +exact (fun v g => Props.In_dec v (V g)). +Qed. + +Definition pmap_delete_preferences v g := +let pm := pmap g in +let m := VertexMap.add v VertexSet.empty pm in +VertexSet.fold + (fun y m' => VertexMap.add y (VertexSet.remove v (adj_set y pm)) m') + (adj_set v pm) + m. + +Lemma delete_preference_sub : forall x v g, +VertexSet.Subset (adj_set x (pmap_delete_preferences v g)) + (adj_set x (pmap g)). + +Proof. +unfold VertexSet.Subset;intros. +unfold pmap_delete_preferences in H. +rewrite VertexSet.fold_1 in H. + +generalize VertexSet.elements_2. +intro H0. generalize (H0 (adj_set v (pmap g))). clear H0. intro H0. +induction (VertexSet.elements (adj_set v (pmap g))); intros. +simpl in H. +unfold adj_set in H. +destruct (Vertex.eq_dec v x). +rewrite InterfFacts.add_eq_o in H. +elim (VertexSet.empty_1 H). +assumption. +rewrite InterfFacts.add_neq_o in H. auto. +assumption. +cut (EqualSetMap + (fold_left + (fun (a : VertexMap.t VertexSet.t) (e : VertexSet.elt) => + VertexMap.add e (VertexSet.remove v (adj_set e (pmap g))) a) (a0 :: l) + (VertexMap.add v VertexSet.empty (pmap g))) + (VertexMap.add a0 + (VertexSet.remove v + (adj_set a0 (pmap g))) + (fold_left + (fun (a : VertexMap.t VertexSet.t) (e : VertexSet.elt) => + VertexMap.add e (VertexSet.remove v (adj_set e (pmap g))) a) l + (VertexMap.add v VertexSet.empty (pmap g))))). + +intro. +assert (VertexSet.In a (adj_set x + (VertexMap.add a0 + (VertexSet.remove v + (adj_set a0 (pmap g))) + (fold_left + (fun (a : VertexMap.t VertexSet.t) (e : VertexSet.elt) => + VertexMap.add e (VertexSet.remove v (adj_set e (pmap g))) a) l + (VertexMap.add v VertexSet.empty (pmap g)))))). +unfold EqualSetMap in H1. generalize (H1 x). intro H2. +inversion H2. +unfold adj_set. unfold adj_set in H5. rewrite <-H5. +unfold adj_set in *. simpl in H. rewrite <-H4 in H. +assumption. +unfold adj_set in *. rewrite <-H4. +unfold adj_set in *. simpl in H. rewrite <-H3 in H. +rewrite <-H5. assumption. +generalize (H1 x). clear H1. intro H1. inversion H1. +simpl in H. unfold adj_set in H. unfold adj_set in H4. rewrite <-H4 in H. +elim (VertexSet.empty_1 H). +clear H1. +destruct (Vertex.eq_dec x a0). +unfold adj_set in H2. rewrite InterfFacts.add_eq_o in H2. +generalize (VertexSet.remove_3 H2). unfold adj_set. +rewrite (InterfFacts.find_o _ e). auto. intuition. +apply IHl. +rewrite InterfFacts.add_neq_o in H4. +unfold adj_set in *. rewrite <-H4. +rewrite <-H5. simpl in H. rewrite <-H3 in *. +assumption. +auto. +intuition. + +apply fold_left_assoc_map. + +intros. +unfold EqualSetMap. intros. +destruct (Vertex.eq_dec x0 z). +rewrite InterfFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y). +rewrite InterfFacts.add_eq_o. +constructor. +generalize (Vertex.eq_trans (Vertex.eq_sym e0) e). intro HH. +unfold adj_set. +rewrite (InterfFacts.find_o _ HH). apply VertexSet.eq_refl. intuition. + +rewrite InterfFacts.add_neq_o. +rewrite InterfFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +auto. +intuition. + +rewrite InterfFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y). +rewrite InterfFacts.add_eq_o. +rewrite InterfFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. + +intros. +unfold EqualSetMap. intros. +destruct (Vertex.eq_dec x0 y). +rewrite InterfFacts.add_eq_o. +rewrite InterfFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. +rewrite InterfFacts.add_neq_o. +rewrite InterfFacts.add_neq_o. +rewrite InterfFacts.add_neq_o. +apply EqualSetMap_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +destruct (Vertex.eq_dec x0 a1). +rewrite InterfFacts.add_eq_o. +rewrite InterfFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. +rewrite InterfFacts.add_neq_o. +rewrite InterfFacts.add_neq_o. +apply (H1 x0). +auto. +auto. +Qed. + +Lemma in_pmap_delete_in : forall x v g, +In_graph v g -> +VertexMap.In x (pmap_delete_preferences v g) -> +VertexMap.In x (pmap g). + +Proof. +intros x v g H1 H. +unfold pmap_delete_preferences in H. +rewrite VertexSet.fold_1 in H. + +generalize VertexSet.elements_2. intro HH. +generalize (HH (adj_set v (pmap g))). clear HH. intro HH. + +induction (VertexSet.elements (adj_set v (pmap g))); intros. +simpl in H. +destruct (proj1 (InterfFacts.add_in_iff _ _ _ _) H). +apply (proj2 (extremities_pmap g x)). +rewrite <-H0. assumption. +assumption. + +cut (EqualSetMap + (fold_left + (fun (a : VertexMap.t VertexSet.t) (e : VertexSet.elt) => + VertexMap.add e (VertexSet.remove v (adj_set e (pmap g))) a) (a :: l) + (VertexMap.add v VertexSet.empty (pmap g))) + (VertexMap.add a + (VertexSet.remove v + (adj_set a (pmap g))) + (fold_left + (fun (a : VertexMap.t VertexSet.t) (e : VertexSet.elt) => + VertexMap.add e (VertexSet.remove v (adj_set e (pmap g))) a) l + (VertexMap.add v VertexSet.empty (pmap g))))). intro. +unfold EqualSetMap in H0. + +generalize (proj1 (InterfFacts.in_find_iff _ _) H). clear H. intro H. +generalize (H0 x). intro H2. +inversion H2. +unfold adj_set in *. simpl in H. rewrite <-H4 in H. +elim H. auto. +clear H2. +destruct (Vertex.eq_dec x a). +rewrite InterfFacts.add_eq_o in H4. +apply (proj2 (InterfFacts.in_find_iff _ _)). +assert (VertexSet.In v (adj_set x (pmap g))). +apply (sym_pmap g). +apply HH. left. auto. +unfold adj_set in H2. +destruct (VertexMap.find x (pmap g)). +intro Helim. inversion Helim. +elim (VertexSet.empty_1 H2). intuition. +rewrite InterfFacts.add_neq_o in H4. +apply IHl. +apply (proj2 (InterfFacts.in_find_iff _ _)). +rewrite <-H4. intro Helim. inversion Helim. +intuition. +auto. + +apply fold_left_assoc_map. +intros. +unfold EqualSetMap. intros. +destruct (Vertex.eq_dec x0 z). +rewrite InterfFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y). +rewrite InterfFacts.add_eq_o. +constructor. +generalize (Vertex.eq_trans (Vertex.eq_sym e0) e). intro HHH. +unfold adj_set. +rewrite (InterfFacts.find_o _ HHH). apply VertexSet.eq_refl. intuition. + +rewrite InterfFacts.add_neq_o. +rewrite InterfFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +auto. +intuition. + +rewrite InterfFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y). +rewrite InterfFacts.add_eq_o. +rewrite InterfFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. + +intros. +unfold EqualSetMap. intros. +destruct (Vertex.eq_dec x0 y). +rewrite InterfFacts.add_eq_o. +rewrite InterfFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. +rewrite InterfFacts.add_neq_o. +rewrite InterfFacts.add_neq_o. +rewrite InterfFacts.add_neq_o. +apply EqualSetMap_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +destruct (Vertex.eq_dec x0 a0). +rewrite InterfFacts.add_eq_o. +rewrite InterfFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. +rewrite InterfFacts.add_neq_o. +rewrite InterfFacts.add_neq_o. +apply (H0 x0). +auto. +auto. +Qed. + +Lemma extremities_in_delete_preferences_imap (v : Vertex.t) g : +forall x, +VertexMap.In x (imap g) <-> VertexSet.In x (V g). + +Proof. +exact (fun v g => extremities_imap g). +Qed. + +Lemma extremities_in_delete_preferences_pmap v g : +In_graph v g -> +(forall x, +VertexMap.In x (pmap_delete_preferences v g) <-> +VertexSet.In x (V g)). + +Proof. +split; intros. +apply (proj1 (extremities_pmap g x)). +apply in_pmap_delete_in with (v:=v); auto. + +unfold pmap_delete_preferences. +set (f := (fun (y : VertexSet.elt) (m' : VertexMap.t VertexSet.t) => + VertexMap.add y (VertexSet.remove v (adj_set y (pmap g))) m')) in *. +set (m := VertexMap.add v VertexSet.empty (pmap g)). +rewrite VertexSet.fold_1. +induction (VertexSet.elements (adj_set v (pmap g))). simpl. +unfold m. +rewrite MapFacts.in_find_iff. +destruct (Vertex.eq_dec x v). +rewrite MapFacts.add_eq_o. +congruence. intuition. +rewrite MapFacts.add_neq_o. +generalize (proj2 (extremities_pmap g x) H0). +rewrite MapFacts.in_find_iff. auto. +auto. +set (h := (fun (a0 : VertexMap.t VertexSet.t) (e : VertexSet.elt) => f e a0)) in *. + +cut (EqualSetMap (fold_left h (a :: l) m) (h (fold_left h l m) a)). intro. +generalize (H1 x). clear H1. intro H1. +inversion H1. +set (tmp := fold_left h l m) in *. +unfold h in H4. +unfold f in H4. +destruct (Vertex.eq_dec x a). +rewrite MapFacts.add_eq_o in H4. +congruence. intuition. +rewrite MapFacts.add_neq_o in H4. +rewrite MapFacts.in_find_iff in IHl. rewrite <-H4 in IHl. +congruence. intuition. +simpl. rewrite MapFacts.in_find_iff. rewrite <-H2. congruence. + +apply fold_left_assoc_map. +unfold EqualSetMap. unfold h. unfold f. intros. +destruct (Vertex.eq_dec x0 z). +rewrite InterfFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y). +rewrite InterfFacts.add_eq_o. +constructor. +generalize (Vertex.eq_trans (Vertex.eq_sym e0) e). intro HHH. +unfold adj_set. +rewrite (InterfFacts.find_o _ HHH). apply VertexSet.eq_refl. +intuition. + +rewrite InterfFacts.add_neq_o. +rewrite InterfFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +auto. +intuition. + +rewrite InterfFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y). +rewrite InterfFacts.add_eq_o. +rewrite InterfFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. + +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s); constructor. +apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +intros. +unfold EqualSetMap. unfold h. unfold f. intros. +destruct (Vertex.eq_dec x0 a0). +rewrite InterfFacts.add_eq_o. +rewrite InterfFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. +rewrite InterfFacts.add_neq_o. +rewrite InterfFacts.add_neq_o. +apply H1. auto. auto. +Qed. + +(* +Lemma extremities_in_delete_preferences v g : +forall x, +VertexMap.In x (imap g) \/ +VertexMap.In x (pmap_delete_preferences v g) <-> +VertexSet.In x (V g). + +Proof. +split; intros. +destruct (Vertex.eq_dec v x). +apply (proj1 (extremities g x)). +destruct H;[left|right]. +assumption. +apply in_pmap_delete_in with (v:=v); auto. + +Qed. +*) + +Lemma simple_graph_delete_preferences v g : forall x y, +VertexSet.In x (adj_set y (imap g)) /\ +VertexSet.In x (adj_set y (pmap_delete_preferences v g)) -> False. + +Proof. +intros. +destruct H. +generalize (delete_preference_sub _ _ _ _ H0). clear H0. intro H0. +apply (simple_graph g x y). intuition. +Qed. + +Lemma sym_imap_delete_preferences (v : Vertex.t) g : +forall x y, +VertexSet.In x (adj_set y (imap g)) -> +VertexSet.In y (adj_set x (imap g)). + +Proof. +exact (fun v g => sym_imap g). +Qed. + +Lemma in_adj_delete_preference_not_eq_1 : forall x y v g, +VertexSet.In x (adj_set y (pmap_delete_preferences v g)) -> +~Vertex.eq y v. + +Proof. +intros. +unfold pmap_delete_preferences in H. +rewrite VertexSet.fold_1 in H. +set (f:= (fun (a : VertexMap.t VertexSet.t) (e : VertexSet.elt) => + VertexMap.add e (VertexSet.remove v (adj_set e (pmap g))) a)) in *. +set (l := VertexSet.elements (adj_set v (pmap g))) in *. +generalize VertexSet.elements_2. intro HH. +generalize (HH (adj_set v (pmap g))). clear HH. intro HH. +fold l in HH. +assert (forall z, In z l -> ~Vertex.eq v z) as Hneq. +clear H. +intros. intro H0. +induction l. inversion H. +simpl in H. destruct H. subst. +assert (VertexSet.In v (adj_set v (pmap g))). +apply HH. left. auto. +elim (not_eq_extremities g v v). +right. assumption. +auto. +apply IHl. +intros. apply HH. auto. +assumption. + +induction l. simpl in H. +unfold adj_set in H. intro Helim. +rewrite MapFacts.add_eq_o in H. +elim (VertexSet.empty_1 H). intuition. +set (s := VertexMap.add v VertexSet.empty (pmap g)) in *. + +assert (EqualSetMap (fold_left f (a :: l) s) (f (fold_left f l s) a)). +apply fold_left_assoc_map. + +unfold f. unfold EqualSetMap. intros. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +constructor. +apply Props.Equal_remove. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_trans (Vertex.eq_sym e) e0)). +apply VertexSet.eq_refl. intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s0); constructor. +apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold f. unfold EqualSetMap. intros. +destruct (Vertex.eq_dec x0 a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H0. +auto. +auto. + +intro Helim. +cut (VertexSet.In x (adj_set y (f (fold_left f l s) a))). +set (tmp := fold_left f l s) in *. +intro H1. +unfold f in H1. +generalize (H0 x). clear H0. intro H0. +unfold adj_set in H1. inversion H0. +rewrite <-H4 in *. +destruct (Vertex.eq_dec y a). +elim (Hneq a). left. auto. +apply (Vertex.eq_trans (Vertex.eq_sym Helim) e). +rewrite MapFacts.add_neq_o in H1. +apply IHl. +assumption. +intros. apply HH. auto. +intros. apply Hneq. right. auto. +assumption. +auto. + +rewrite <-H3 in *. +destruct (Vertex.eq_dec y a). +elim (Hneq a). left. auto. +apply (Vertex.eq_trans (Vertex.eq_sym Helim) e). +rewrite MapFacts.add_neq_o in H1. +apply IHl. +assumption. +intros. apply HH. auto. +intros. apply Hneq. right. auto. +assumption. +auto. + +clear HH Hneq IHl. +unfold adj_set in H. unfold adj_set. +generalize (H0 y). clear H0. intro H0. +simpl in H. +inversion H0; subst. +rewrite <-H3 in *. +simpl in H. rewrite <-H2 in *. +elim (VertexSet.empty_1 H). +rewrite <-H1 in *. +rewrite <-H2 in *. +rewrite <-H3. assumption. +Qed. + +Lemma in_adj_delete_preference_not_eq_2 : forall x y v g, +VertexSet.In x (adj_set y (pmap_delete_preferences v g)) -> +~Vertex.eq x v. + +Proof. +intros. intro Helim. generalize H. intro Hcopy. +unfold pmap_delete_preferences in H. +rewrite VertexSet.fold_1 in H. +set (f:= (fun (a : VertexMap.t VertexSet.t) (e : VertexSet.elt) => + VertexMap.add e (VertexSet.remove v (adj_set e (pmap g))) a)) in *. +generalize VertexSet.elements_1. intro. +generalize (H0 (adj_set v (pmap g)) y). clear H0. intro HH. +set (l := VertexSet.elements (adj_set v (pmap g))) in *. +induction l. simpl in H. +unfold adj_set in H. +rewrite MapFacts.add_neq_o in H. +generalize (sym_pmap g). intro Hsym. +generalize (Hsym _ _ H). intro H0. +unfold adj_set in H0. +rewrite (MapFacts.find_o _ Helim) in H0. +generalize (HH H0). intro H1. inversion H1. +generalize (in_adj_delete_preference_not_eq_1 x y v g Hcopy). auto. +set (s := VertexMap.add v VertexSet.empty (pmap g)) in *. + +assert (EqualSetMap (fold_left f (a :: l) s) (f (fold_left f l s) a)). +apply fold_left_assoc_map. + +unfold f. unfold EqualSetMap. intros. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +constructor. +apply Props.Equal_remove. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_trans (Vertex.eq_sym e) e0)). +apply VertexSet.eq_refl. intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s0); constructor. +apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold f. unfold EqualSetMap. intros. +destruct (Vertex.eq_dec x0 a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H0. +auto. +auto. + +clear Hcopy. +cut (VertexSet.In x (adj_set y (f (fold_left f l s) a))). +set (tmp := fold_left f l s) in *. +intro H1. +unfold f in H1. unfold adj_set in H1. +destruct (Vertex.eq_dec y a). +rewrite MapFacts.add_eq_o in H1. +elim (VertexSet.remove_1 (Vertex.eq_sym Helim) H1). intuition. +rewrite MapFacts.add_neq_o in H1. +apply IHl. +assumption. +intro. generalize (HH H2). intro H3. +inversion H3; subst. +elim (n H5). +assumption. +auto. + +unfold adj_set. +generalize (H0 y). clear H0. intro H0. +inversion H0; subst. +simpl in H. unfold adj_set in H. rewrite <-H2 in H. +assumption. +simpl in H. unfold adj_set in H. rewrite <-H1 in H. +rewrite <-H3. assumption. +Qed. + +Lemma sym_pmap_delete_preferences (v : Vertex.t) g : +forall x y, +VertexSet.In x (adj_set y (pmap_delete_preferences v g)) -> +VertexSet.In y (adj_set x (pmap_delete_preferences v g)). + +Proof. +intros. +generalize (in_adj_delete_preference_not_eq_1 x y v g H). intro Hneqy. +generalize (in_adj_delete_preference_not_eq_2 x y v g H). intro Hneqx. +generalize (delete_preference_sub y v g x H). intro HH. +generalize (sym_pmap g _ _ HH). clear HH. intro HH. +unfold pmap_delete_preferences in *. +rewrite VertexSet.fold_1 in *. +set (f := fun (a : VertexMap.t VertexSet.t) (e : VertexSet.elt) => + VertexMap.add e (VertexSet.remove v (adj_set e (pmap g))) a) in *. +induction (VertexSet.elements (adj_set v (pmap g))). +simpl in *. +destruct (Vertex.eq_dec x v). +destruct (Vertex.eq_dec y v). +unfold adj_set in H. rewrite MapFacts.add_eq_o in H. +elim (VertexSet.empty_1 H). +case_eq (VertexMap.find y (VertexMap.add v VertexSet.empty (pmap g))). intros. +rewrite H0 in H. +rewrite MapFacts.add_eq_o in H0. inversion H0. subst. +elim (VertexSet.empty_1 H). intuition. +rewrite MapFacts.add_eq_o. congruence. intuition. + +elim (Hneqx e). + +unfold adj_set. rewrite MapFacts.add_neq_o. assumption. intuition. + +set (s := VertexMap.add v VertexSet.empty (pmap g)) in *. +assert (EqualSetMap (fold_left f (a :: l) s) (f (fold_left f l s) a)). +apply fold_left_assoc_map. + +unfold f. unfold EqualSetMap. intros. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +constructor. +apply Props.Equal_remove. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_trans (Vertex.eq_sym e) e0)). +apply VertexSet.eq_refl. intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s0); constructor. +apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold f. unfold EqualSetMap. intros. +destruct (Vertex.eq_dec x0 a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H0. +auto. +auto. + +unfold EqualSetMap in H0. +unfold adj_set. +case_eq (VertexMap.find x (fold_left f (a :: l) s)); intros. +generalize (H0 x). intro HH0. +rewrite H1 in HH0. inversion HH0. subst. +set (tmp := fold_left f l s) in *. +unfold adj_set in H. +unfold f in H3. +destruct (Vertex.eq_dec x a). +rewrite MapFacts.add_eq_o in H3. +inversion H3. subst. clear H3. +rewrite H4. apply VertexSet.remove_2. +auto. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e)). assumption. +intuition. + +rewrite MapFacts.add_neq_o in H3. +rewrite H4. +cut (VertexSet.In y (adj_set x tmp)). +intro. +unfold adj_set in H2. +rewrite <-H3 in H2. assumption. +apply IHl. +assert (VertexSet.In x (adj_set y (pmap g)) -> VertexSet.In x (adj_set y tmp)). +clear IHl H0 H1 HH0 H4 H3 HH H. +intros. +unfold tmp. +induction l. simpl. unfold s. +unfold adj_set. +rewrite MapFacts.add_neq_o. assumption. +auto. + +assert (EqualSetMap (fold_left f (a0 :: l) s) (f (fold_left f l s) a0)). +apply fold_left_assoc_map. + +unfold f. unfold EqualSetMap. intros. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +constructor. +apply Props.Equal_remove. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_trans (Vertex.eq_sym e) e0)). +apply VertexSet.eq_refl. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s0); constructor. +apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold f. unfold EqualSetMap. intros. +destruct (Vertex.eq_dec x0 a1). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H0. +auto. +auto. + +generalize (H0 y). clear H0. intro H0. +inversion H0; subst. +unfold adj_set. simpl. rewrite <-H2. +set (tmp' := fold_left f l s) in *. +unfold f in H3. +destruct (Vertex.eq_dec y a0). +rewrite MapFacts.add_eq_o in H3. inversion H3. +intuition. +rewrite MapFacts.add_neq_o in H3. +unfold adj_set in IHl. rewrite <-H3 in IHl. +assumption. +auto. +unfold adj_set. simpl. rewrite <-H1. +rewrite H3. +set (tmp' := fold_left f l s) in *. +unfold f in H2. +destruct (Vertex.eq_dec y a0). +rewrite MapFacts.add_eq_o in H2. +inversion H2. subst. clear H2. +apply VertexSet.remove_2. +auto. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e)). +assumption. +intuition. +rewrite MapFacts.add_neq_o in H2. +assert (VertexSet.In x (adj_set y tmp')). +apply IHl. +unfold adj_set in H4. +rewrite <-H2 in H4. assumption. +auto. +apply H2. +apply (sym_pmap g). assumption. +auto. + +generalize (H0 x). clear H0. intro H0. +rewrite H1 in H0. inversion H0. subst. +set (tmp := fold_left f l s) in *. +unfold f in H2. +destruct (Vertex.eq_dec x a). +rewrite MapFacts.add_eq_o in H2. +inversion H2. +intuition. +rewrite MapFacts.add_neq_o in H2. + +assert (VertexSet.In y (adj_set x (pmap g)) -> VertexSet.In y (adj_set x tmp)). +clear IHl H H0 HH H2 H1. +intros. +unfold tmp. +induction l. simpl. unfold s. +unfold adj_set. +rewrite MapFacts.add_neq_o. assumption. +intuition. + +assert (EqualSetMap (fold_left f (a0 :: l) s) (f (fold_left f l s) a0)). +apply fold_left_assoc_map. + +unfold f. unfold EqualSetMap. intros. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +constructor. +apply Props.Equal_remove. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_trans (Vertex.eq_sym e) e0)). +apply VertexSet.eq_refl. intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s0); constructor. +apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold f. unfold EqualSetMap. intros. +destruct (Vertex.eq_dec x0 a1). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H0. +auto. +auto. + +generalize (H0 x). clear H0. intro H0. +inversion H0; subst. +unfold adj_set. simpl. rewrite <-H2. +set (tmp' := fold_left f l s) in *. +unfold f in H3. +destruct (Vertex.eq_dec x a0). +rewrite MapFacts.add_eq_o in H3. inversion H3. +intuition. +rewrite MapFacts.add_neq_o in H3. +unfold adj_set in IHl. rewrite <-H3 in IHl. +assumption. +auto. +unfold adj_set. simpl. rewrite <-H1. +rewrite H3. +set (tmp' := fold_left f l s) in *. +unfold f in H2. +destruct (Vertex.eq_dec x a0). +rewrite MapFacts.add_eq_o in H2. +inversion H2. subst. clear H2. +apply VertexSet.remove_2. +auto. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e)). +assumption. +intuition. +rewrite MapFacts.add_neq_o in H2. +assert (VertexSet.In y (adj_set x tmp')). +apply IHl. +unfold adj_set in H4. +rewrite <-H2 in H4. assumption. +auto. +assert (VertexSet.In y (adj_set x tmp)). +apply H3. assumption. +unfold adj_set in H4. +rewrite <-H2 in H4. +assumption. +auto. +Qed. + +Lemma not_eq_extremities_delete_preferences v g : +forall x y, +VertexSet.In x (adj_set y (imap g)) \/ +VertexSet.In x (adj_set y (pmap_delete_preferences v g)) -> +~Vertex.eq x y. + +Proof. +intros. +apply (not_eq_extremities g). +destruct H;[left|right]. +assumption. +apply (delete_preference_sub _ _ _ _ H). +Qed. + +Definition delete_preference_edges v g H := +Make_Graph (V g) + (imap g) + (pmap_delete_preferences v g) + (extremities_in_delete_preferences_imap v g) + (extremities_in_delete_preferences_pmap v g H) + (simple_graph_delete_preferences v g) + (sym_imap_delete_preferences v g) + (sym_pmap_delete_preferences v g) + (not_eq_extremities_delete_preferences v g). + +Definition Prefere x y g := +exists w, In_graph_edge (x,y,Some w) g. + +Definition Interfere x y g := +In_graph_edge (x,y,None) g. + +Definition precolored g := +VertexSet.filter (fun v => is_precolored v g) (V g). + +Parameter mreg_ext : forall x y, +Vertex.eq x y -> is_mreg x = is_mreg y. + +Lemma mem_ext : forall x y s, +Vertex.eq x y -> VertexSet.mem x s = VertexSet.mem y s. + +Proof. +intros. +case_eq (VertexSet.mem x s); intros. +generalize (VertexSet.mem_2 H0). intro H1. +rewrite H in H1. generalize (VertexSet.mem_1 H1). auto. +case_eq (VertexSet.mem y s); intros. +generalize (VertexSet.mem_2 H1). intro H2. +rewrite <-H in H2. generalize (VertexSet.mem_1 H2). intro H3. +rewrite H3 in H0. inversion H0. +reflexivity. +Qed. + +Lemma compat_bool_is_precolored : forall g, +compat_bool Vertex.eq (fun x => is_precolored x g). + +Proof. +unfold compat_bool. intros. +unfold is_precolored. +apply mreg_ext. auto. +Qed. + +(* Equivalence of predicate and function *) +Lemma precolored_equiv : forall x g, +VertexSet.In x (precolored g) <-> (is_precolored x g = true /\ In_graph x g). + +Proof. +split; intros. +split. +apply (VertexSet.filter_2 (compat_bool_is_precolored _) H). +apply (VertexSet.filter_1 (compat_bool_is_precolored _) H). +apply VertexSet.filter_3. apply compat_bool_is_precolored. +unfold is_precolored in H. +case_eq (VertexSet.mem x (V g)); intros. +apply VertexSet.mem_2. assumption. intuition. intuition. +Qed. + +Lemma In_graph_aff_edge_in_AE : forall e g, +EdgeSet.In e (AE g) <-> aff_edge e /\ In_graph_edge e g. + +Proof. +intros e g. split; intro H. +split. +unfold aff_edge. exists N0. exact (AE_weights g e H). +left. assumption. +destruct H as [H H0]. +destruct H0. +assumption. +unfold aff_edge in H. destruct H as [w H]. +rewrite (IE_weights g e H0) in H. inversion H. +Qed. + +Lemma In_graph_interf_edge_in_IE : forall e g, +EdgeSet.In e (IE g) <-> interf_edge e /\ In_graph_edge e g. + +Proof. +intros e g. split; intro H. +split. +exact (IE_weights g e H). +right. assumption. +destruct H as [H H0]. +destruct H0. +unfold interf_edge in H. +rewrite (AE_weights g e H0) in H. inversion H. +assumption. +Qed. + +(* Spec of Prefere *) +Lemma Prefere_1 : forall e g, +aff_edge e -> +In_graph_edge e g -> +Prefere (fst_ext e) (snd_ext e) g. + +Proof. +intros. +unfold Prefere. +destruct H0. +exists N0. rewrite <-(AE_weights g e H0). rewrite <-edge_eq. left. assumption. +generalize (proj1 (In_graph_interf_edge_in_IE _ _) H0). intro H1. +destruct H1 as [H1 _]. unfold interf_edge in H1. unfold aff_edge in H. +rewrite H1 in H. inversion H. inversion H2. +Qed. + +Lemma Prefere_2 : forall x y g, +Prefere x y g -> +In_graph_edge (x,y,Some N0) g. + +Proof. +intros. +unfold Prefere in H. +destruct H. +assert (EdgeSet.In (x,y,Some x0) (AE g)). +rewrite In_graph_aff_edge_in_AE. split. exists x0. auto. auto. +generalize (AE_weights _ _ H0). intro. simpl in H1. rewrite H1 in H. auto. +Qed. + +(* spec of Interfere *) +Lemma Interfere_1 : forall e g, +interf_edge e -> +In_graph_edge e g -> +Interfere (fst_ext e) (snd_ext e) g. + +Proof. +intros. +unfold Interfere. +destruct H0. +generalize (proj1 (In_graph_aff_edge_in_AE _ _) H0). intro H1. +destruct H1 as [H1 _]. unfold interf_edge in H. unfold aff_edge in H1. +rewrite H in H1. inversion H1. inversion H2. +rewrite <-(IE_weights g e H0). rewrite <-edge_eq. right. assumption. +Qed. + +Lemma Interfere_2 : forall x y g, +Interfere x y g -> +In_graph_edge (x,y,None) g. + +Proof. +auto. +Qed. + +(* Machine registers are registers *) +Lemma In_precolored : forall g , +VertexSet.Subset (precolored g) (V g). + +Proof. +intros. +unfold VertexSet.Subset; intros. +apply (VertexSet.filter_1 (compat_bool_is_precolored _) H). +Qed. + +(* specification of remove_vertex *) +Lemma In_remove : forall x r g, +In_graph x g -> +~Vertex.eq x r -> +In_graph x (remove_vertex r g). + +Proof. +intros. +unfold remove_vertex. simpl. +apply VertexSet.remove_2; auto. +Qed. + +(* A precolored vertex cannot be removed from the graph *) +Lemma not_in_remove : forall r g, +~In_graph r (remove_vertex r g). + +Proof. +intros. +unfold remove_vertex. +apply VertexSet.remove_1. apply Vertex.eq_refl. +Qed. + +Lemma in_remove_in : forall x r g, +In_graph x (remove_vertex r g) -> +In_graph x g. + +Proof. +intros. +unfold remove_vertex in H. +unfold In_graph in H. simpl in H. +apply (VertexSet.remove_3 H). +Qed. + +Add Morphism In_graph : In_graph_m. + +Proof. +intros. +unfold In_graph. +rewrite H. reflexivity. +Qed. + +(* Probably redundant, TODO get out of interface *) +Lemma precolored_remove_vertex2 : forall x y g, +VertexSet.In x (VertexSet.remove y (precolored g)) <-> +VertexSet.In x (precolored (remove_vertex y g)). + +Proof. +intros. +split; intros. +generalize (VertexSet.remove_3 H). intro H0. +generalize (VertexSet.filter_1 (compat_bool_is_precolored _) H0). intro. +generalize (VertexSet.filter_2 (compat_bool_is_precolored _) H0). clear H0. intro. +apply VertexSet.filter_3. +apply compat_bool_is_precolored. +unfold remove_vertex. simpl. +apply VertexSet.remove_2. intro. apply (VertexSet.remove_1 H2 H). +assumption. +unfold is_precolored. assumption. +generalize (proj1 (precolored_equiv _ _) H). clear H. intro H. destruct H. +apply VertexSet.remove_2. +intro H1. rewrite <-H1 in H0. elim (not_in_remove _ _ H0). +apply (proj2 (precolored_equiv _ _)). +split. +assumption. +apply in_remove_in with (r:=y). assumption. +Qed. + +Lemma In_remove_edge_ : forall e r g, +In_graph_edge e g -> +~incident e r -> +In_graph_edge e (remove_vertex r g). + +Proof. +intros. +destruct H;[left|right]. +generalize (AE_weights _ _ H). intro Hw. +unfold AE in *. +unfold remove_vertex. +simpl. +rewrite (edge_eq e) in *. +simpl in Hw. rewrite Hw. +apply (proj2 (edgemap_to_edgeset_charac _ _ _ _ (sym_pmap_remove_vertex r g))). +apply imap_remove_1. +intro H1. elim H0. right. auto. +intro H1. elim H0. left. auto. +rewrite Hw in H. +apply (proj1 (edgemap_to_edgeset_charac _ _ _ _ (sym_pmap g)) H). + +generalize (IE_weights _ _ H). intro Hw. +unfold IE in *. +unfold remove_vertex. +simpl. +rewrite (edge_eq e) in *. +simpl in Hw. rewrite Hw. +apply (proj2 (edgemap_to_edgeset_charac _ _ _ _ (sym_imap_remove_vertex r g))). +apply imap_remove_1. +intro H1. elim H0. right. auto. +intro H1. elim H0. left. auto. +rewrite Hw in H. +apply (proj1 (edgemap_to_edgeset_charac _ _ _ _ (sym_imap g)) H). +Qed. + +Lemma in_remove_in_edge : forall e r g, +In_graph_edge e (remove_vertex r g) -> +In_graph_edge e g. + +Proof. +intros. +destruct H;[left|right]. +generalize (AE_weights _ _ H). intro Hw. +unfold AE in *. +unfold remove_vertex in H. +simpl in H. +rewrite (edge_eq e) in *. +simpl in Hw. rewrite Hw. +apply (proj2 (edgemap_to_edgeset_charac _ _ _ _ (sym_pmap g))). +apply imap_remove_3 with (r:=r). +rewrite Hw in H. +apply (proj1 (edgemap_to_edgeset_charac _ _ _ _(sym_pmap_remove_vertex r g)) H). + +generalize (IE_weights _ _ H). intro Hw. +unfold IE in *. +unfold remove_vertex in H. +simpl in H. +rewrite (edge_eq e) in *. +simpl in Hw. rewrite Hw. +apply (proj2 (edgemap_to_edgeset_charac _ _ _ _ (sym_imap g))). +apply imap_remove_3 with (r:=r). +rewrite Hw in H. +apply (proj1 (edgemap_to_edgeset_charac _ _ _ _(sym_imap_remove_vertex r g)) H). +Qed. + +Lemma not_in_remove_edge : forall e r g, +incident e r -> +~In_graph_edge e (remove_vertex r g). + +Proof. +intros. generalize H. intro H0. +destruct (In_graph_edge_dec e g). +assert (In_graph r g) as Hin. +destruct H. +rewrite H. apply (proj1 (In_graph_edge_in_ext _ _ i)). +rewrite H. apply (proj2 (In_graph_edge_in_ext _ _ i)). +intro H1. destruct H1. +generalize (AE_weights _ _ H1). intro Hw. +unfold AE in *. +unfold remove_vertex in H1. +simpl in H1. +rewrite (edge_eq e) in H1. +rewrite Hw in H1. rewrite edge_comm in H1. +generalize (proj1 (edgemap_to_edgeset_charac _ _ _ _ (sym_pmap_remove_vertex r g)) H1). +intro H3. +apply (imap_remove_2 _ _ (pmap g) _ H (sym_pmap g) H3). + +generalize (IE_weights _ _ H1). intro Hw. +unfold IE in *. +unfold remove_vertex in H1. +simpl in H1. +rewrite (edge_eq e) in H1. +rewrite Hw in H1. rewrite edge_comm in H1. +generalize (proj1 (edgemap_to_edgeset_charac _ _ _ _ (sym_imap_remove_vertex r g)) H1). +intro H3. +apply (imap_remove_2 _ _ (imap g) _ H (sym_imap g) H3). +intro H1. elim (n (in_remove_in_edge _ _ _ H1)). +Qed. + + +(* spec of merge *) +Lemma In_merge_in : forall g e x p q, +In_graph x (merge e g p q) -> In_graph x g. + +Proof. +intros. +unfold merge in H. +destruct (aff_edge_dec e). +destruct (In_graph_edge_dec e g). +unfold In_graph in H. simpl in H. +apply (VertexSet.remove_3 H). +elim (n p). +elim (n q). +Qed. + +(* the second extremity must not be precolored, + cause it is removed from the graph *) +Lemma merge_1 : forall e g p q, +~In_graph (snd_ext e) (merge e g p q). + +Proof. +intros. +unfold merge. unfold In_graph. simpl. +apply VertexSet.remove_1. apply Vertex.eq_refl. +Qed. + +Lemma merge_2 : forall x e g p q, +~Vertex.eq x (snd_ext e) -> +In_graph x g -> +In_graph x (merge e g p q). + +Proof. +intros. unfold merge. +unfold In_graph. simpl. apply VertexSet.remove_2 with (x:=(snd_ext e)); auto. +Qed. + +Lemma merge_3 : forall e e' g p q, +In_graph_edge e' g -> +interf_edge e' -> +In_graph_edge (redirect (snd_ext e) (fst_ext e) e') (merge e g p q). + +Proof. +intros e e' g p q H0 H1. generalize I. intro H. +assert (EdgeSet.In (fst_ext e, snd_ext e, Some N0) (AE g)) as He. +destruct p. rewrite (edge_eq e) in H2. +generalize (AE_weights _ _ H2). intro. simpl in H3. +rewrite H3 in H2. assumption. +generalize (IE_weights _ _ H2). inversion q. congruence. +assert (VertexSet.In (snd_ext e) (adj_set (fst_ext e) (pmap g))) as Hee. +apply (proj1 (edgemap_to_edgeset_charac _ _ _ (Some N0) (sym_pmap g))). +assumption. +assert (EdgeSet.In (fst_ext e', snd_ext e', None) (IE g)) as He'. +destruct H0. generalize (AE_weights _ _ H0). intro H2. inversion H2. congruence. +rewrite (edge_eq e') in H0. generalize (IE_weights _ _ H0). intro. simpl in H2. +rewrite H2 in H0. assumption. +assert (VertexSet.In (snd_ext e') (adj_set (fst_ext e') (imap g))) as Hee'. +apply (proj1 (edgemap_to_edgeset_charac _ _ _ None (sym_imap g))). +assumption. +assert (forall a b, VertexSet.In a (adj_set b (imap_merge e g)) -> + VertexSet.In b (adj_set a (imap_merge e g))) as Hsym. +apply sym_imap_merge_map; auto. + +right. +unfold merge. unfold IE. simpl. +unfold redirect. +destruct (OTFacts.eq_dec (fst_ext e') (snd_ext e)). +rewrite H1. apply (proj2 (edgemap_to_edgeset_charac (imap_merge e g) _ _ None Hsym)). + +unfold imap_merge. +unfold map_merge. +set (s := (VertexSet.union + (VertexSet.remove (snd_ext e) (adj_set (fst_ext e) (imap g))) + (VertexSet.remove (fst_ext e) (adj_set (snd_ext e) (imap g))))) in *. +set (m := imap g) in *. +set (f := (fun (y : VertexSet.elt) (m' : VertexMap.t VertexSet.t) => + VertexMap.add y + (VertexSet.add (fst_ext e) + (VertexSet.remove (snd_ext e) (adj_set y m))) m')) in *. +set (s' := VertexSet.remove (fst_ext e) (adj_set (snd_ext e) m)) in *. +unfold adj_set. +rewrite MapFacts.remove_neq_o. +rewrite MapFacts.add_eq_o. +unfold s. +apply VertexSet.union_3. +unfold s'. +apply VertexSet.remove_2. +intro. +elim (simple_graph g (fst_ext e) (snd_ext e)). +split. +rewrite H2. unfold adj_set. rewrite (MapFacts.find_o _ (Vertex.eq_sym r)). +fold (adj_set (fst_ext e') (imap g)). +assumption. +apply (sym_pmap g). assumption. +unfold adj_set. rewrite (MapFacts.find_o _ (Vertex.eq_sym r)). +fold (adj_set (fst_ext e') m). +assumption. +intuition. +apply (In_graph_edge_diff_ext _ _ p). + +destruct (OTFacts.eq_dec (snd_ext e') (snd_ext e)). +rewrite H1. +apply (proj2 (edgemap_to_edgeset_charac (imap_merge e g) _ _ None Hsym)). +unfold imap_merge. +unfold map_merge. +set (s := (VertexSet.union + (VertexSet.remove (snd_ext e) (adj_set (fst_ext e) (imap g))) + (VertexSet.remove (fst_ext e) (adj_set (snd_ext e) (imap g))))) in *. +set (m := imap g) in *. +set (f := (fun (y : VertexSet.elt) (m' : VertexMap.t VertexSet.t) => + VertexMap.add y + (VertexSet.add (fst_ext e) + (VertexSet.remove (snd_ext e) (adj_set y m))) m')) in *. +set (s' := VertexSet.remove (fst_ext e) (adj_set (snd_ext e) m)) in *. +unfold adj_set. +rewrite MapFacts.remove_neq_o. +rewrite MapFacts.add_neq_o. +rewrite VertexSet.fold_1. +set (f' := fun a e => f e a). +generalize VertexSet.elements_1. intro HH. +generalize (HH s' (fst_ext e')). clear HH. intro HH. +induction (VertexSet.elements s'). simpl. +assert (InA Vertex.eq (fst_ext e') nil). +apply HH. +unfold s'. apply VertexSet.remove_2. +intro. +elim (simple_graph g (fst_ext e) (snd_ext e)). +split. +rewrite H2. unfold adj_set. rewrite (MapFacts.find_o _ (Vertex.eq_sym r)). +fold (adj_set (fst_ext e') (imap g)). +apply (sym_imap g). assumption. +apply (sym_pmap g). assumption. +unfold adj_set. rewrite (MapFacts.find_o _ (Vertex.eq_sym r)). +apply (sym_imap g). assumption. +inversion H2. + +cut (EqualSetMap (fold_left f' (a :: l) m) (f' (fold_left f' l m) a)). intro. +generalize (H2 (fst_ext e')). clear H2. intro H2. simpl. inversion H2;clear H2. +set (tmp := fold_left f' l m) in *. +unfold f' in H5. unfold f in H5. +destruct (Vertex.eq_dec (fst_ext e') a). +rewrite MapFacts.add_eq_o in H5. congruence. intuition. +rewrite MapFacts.add_neq_o in H5. rewrite <-H5 in IHl. +apply IHl. +intro. generalize (HH H2). clear HH. intro HH. +inversion HH; subst. +elim (n0 H6). +assumption. +auto. + +set (tmp := fold_left f' l m) in *. +unfold f' in H4. unfold f in H4. +destruct (Vertex.eq_dec (fst_ext e') a). +rewrite MapFacts.add_eq_o in H4. +rewrite H5. inversion H4; subst; clear H4. +apply VertexSet.add_1. intuition. intuition. +rewrite MapFacts.add_neq_o in H4. +rewrite <-H4 in IHl. +rewrite H5. apply IHl. +intro. generalize (HH H2). clear HH. intro HH. inversion HH; subst. +elim (n0 H7). +assumption. +auto. + +apply fold_left_assoc_map. +unfold EqualSetMap. unfold f'. unfold f. +intros. +destruct (Vertex.eq_dec x z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x y). +rewrite MapFacts.add_eq_o. +constructor. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). +rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +apply VertexSet.eq_refl. intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x y). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x s0); constructor; apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f'. unfold f. +destruct (Vertex.eq_dec x a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H2. auto. auto. + +intro. +elim (simple_graph g (fst_ext e) (snd_ext e)). +split. +rewrite H2. unfold adj_set. rewrite (MapFacts.find_o _ (Vertex.eq_sym r)). +fold (adj_set (fst_ext e') (imap g)). +apply (sym_imap g). assumption. +apply (sym_pmap g). assumption. + +intuition. +rewrite (edge_eq e'). rewrite H1. +apply (proj2 (edgemap_to_edgeset_charac (imap_merge e g) _ _ None Hsym)). +unfold imap_merge. +unfold map_merge. +set (s := (VertexSet.union + (VertexSet.remove (snd_ext e) (adj_set (fst_ext e) (imap g))) + (VertexSet.remove (fst_ext e) (adj_set (snd_ext e) (imap g))))) in *. +set (m := imap g) in *. +set (f := (fun (y : VertexSet.elt) (m' : VertexMap.t VertexSet.t) => + VertexMap.add y + (VertexSet.add (fst_ext e) + (VertexSet.remove (snd_ext e) (adj_set y m))) m')) in *. +set (s' := VertexSet.remove (fst_ext e) (adj_set (snd_ext e) m)) in *. +unfold adj_set. +rewrite MapFacts.remove_neq_o. +destruct (Vertex.eq_dec (fst_ext e') (fst_ext e)). +rewrite MapFacts.add_eq_o. +unfold s. +apply VertexSet.union_2. +apply VertexSet.remove_2. intuition. +unfold adj_set. rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +assumption. +intuition. + +rewrite MapFacts.add_neq_o. +rewrite VertexSet.fold_1. +set (f' := fun a e => f e a). +(* +generalize VertexSet.elements_1. intro HH. +generalize (HH s' (fst_ext e')). clear HH. intro HH. +*) +induction (VertexSet.elements s'). simpl. assumption. +cut (EqualSetMap (fold_left f' (a :: l) m) (f' (fold_left f' l m) a)). intro. +generalize (H2 (fst_ext e')). clear H2. intro H2. simpl. inversion H2;clear H2. +set (tmp := fold_left f' l m) in *. +unfold f' in H5. unfold f in H5. +destruct (Vertex.eq_dec (fst_ext e') a). +rewrite MapFacts.add_eq_o in H5. congruence. intuition. +rewrite MapFacts.add_neq_o in H5. rewrite <-H5 in IHl. +apply IHl. auto. + +set (tmp := fold_left f' l m) in *. +unfold f' in H4. unfold f in H4. +destruct (Vertex.eq_dec (fst_ext e') a). +rewrite MapFacts.add_eq_o in H4. +rewrite H5. inversion H4; subst; clear H4. +apply VertexSet.add_2. apply VertexSet.remove_2. intuition. +unfold adj_set. rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). assumption. +intuition. + +rewrite MapFacts.add_neq_o in H4. +rewrite <-H4 in IHl. +rewrite H5. apply IHl. +auto. + +apply fold_left_assoc_map. +unfold EqualSetMap. unfold f'. unfold f. +intros. +destruct (Vertex.eq_dec x z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x y). +rewrite MapFacts.add_eq_o. +constructor. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). +rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +apply VertexSet.eq_refl. intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x y). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x s0); constructor; apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f'. unfold f. +destruct (Vertex.eq_dec x a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H2. auto. auto. + +auto. +intuition. +Qed. + +Lemma merge_4 : forall e e' g p q, +In_graph_edge e' (merge e g p q) -> +exists a, In_graph_edge a g /\ weak_eq e' (redirect (snd_ext e) (fst_ext e) a) + /\ same_type a e'. + +Proof. +intros e e' g p q H0. generalize I. intro H. +generalize H0. intro Hin. +unfold merge in H0. destruct H0. +generalize (AE_weights _ _ H0). intro Hw. +unfold AE in H0. simpl in H0. +rewrite (edge_eq e') in H0. rewrite Hw in H0. +generalize (proj1 (edgemap_to_edgeset_charac _ _ _ (Some N0) (sym_pmap_merge_map _ _ q p)) H0). +clear H0. intro H0. +generalize (pmap_merge_sub _ _ _ _ H0). clear H0. intro H0. +unfold map_merge in H0. +set (s := (VertexSet.union + (VertexSet.remove (snd_ext e) (adj_set (fst_ext e) (pmap g))) + (VertexSet.remove (fst_ext e) (adj_set (snd_ext e) (pmap g))))) in *. +set (m := pmap g) in *. +set (f := (fun (y : VertexSet.elt) (m' : VertexMap.t VertexSet.t) => + VertexMap.add y + (VertexSet.add (fst_ext e) + (VertexSet.remove (snd_ext e) (adj_set y m))) m')) in *. +set (s' := VertexSet.remove (fst_ext e) (adj_set (snd_ext e) m)) in *. +unfold adj_set in H0. +destruct (Vertex.eq_dec (fst_ext e') (snd_ext e)). +rewrite MapFacts.remove_eq_o in H0. +elim (VertexSet.empty_1 H0). intuition. +rewrite MapFacts.remove_neq_o in H0. +destruct (Vertex.eq_dec (fst_ext e') (fst_ext e)). +rewrite MapFacts.add_eq_o in H0. +unfold s in H0. +destruct (VertexSet.union_1 H0). +generalize (VertexSet.remove_3 H1). clear H1. intro H1. +exists (fst_ext e', snd_ext e', Some N0). split. +left. unfold AE. +apply (proj2 (edgemap_to_edgeset_charac _ _ _ (Some N0) (sym_pmap g))). +unfold adj_set. rewrite (MapFacts.find_o _ e0). assumption. +unfold redirect. change_rewrite. +destruct (OTFacts.eq_dec (fst_ext e') (snd_ext e)). +elim (n r). +destruct (OTFacts.eq_dec (snd_ext e') (snd_ext e)). +elim (merge_1 e g p q). +rewrite <-r. apply (proj2 (In_graph_edge_in_ext _ _ Hin)). +split. +unfold weak_eq. change_rewrite. +left. split; auto. +unfold same_type. left. split. exists N0. auto. exists N0. auto. +unfold s' in H1. +generalize (VertexSet.remove_3 H1). intro H2. +exists (snd_ext e', snd_ext e, Some N0). +split. +left. unfold AE. +apply (proj2 (edgemap_to_edgeset_charac _ _ _ (Some N0) (sym_pmap g))). +apply (sym_pmap g). assumption. +unfold redirect. change_rewrite. +destruct (OTFacts.eq_dec (snd_ext e') (snd_ext e)). +unfold weak_eq. change_rewrite. +split. +left. split; auto. +unfold same_type. left. split; exists N0; auto. +destruct (OTFacts.eq_dec (snd_ext e) (snd_ext e)). +unfold weak_eq. change_rewrite. +split. +right. split; auto. +unfold same_type. left. split; exists N0; auto. +elim n1. intuition. intuition. +rewrite MapFacts.add_neq_o in H0. + +rewrite VertexSet.fold_1 in H0. +set (f' := fun a e => f e a) in *. +generalize VertexSet.elements_2. intro HH. +generalize (HH s'). clear HH. intro HH. +induction (VertexSet.elements s'). simpl in H0. +case_eq (VertexMap.find (fst_ext e') m); intros. rewrite H1 in H0. +exists (fst_ext e', snd_ext e', Some N0). +split. +left. apply (proj2 (edgemap_to_edgeset_charac _ _ _ (Some N0) (sym_pmap g))). +unfold adj_set. fold m. rewrite H1. assumption. +unfold redirect. change_rewrite. +destruct (OTFacts.eq_dec (fst_ext e') (snd_ext e)). +elim (n r). +destruct (OTFacts.eq_dec (snd_ext e') (snd_ext e)). +elim (merge_1 e g p q). +rewrite <-r. apply (proj2 (In_graph_edge_in_ext _ _ Hin)). +unfold weak_eq. change_rewrite. +split. +left. split ; auto. +left. split; exists N0; auto. +rewrite H1 in H0. elim (VertexSet.empty_1 H0). +simpl in H0. +cut (EqualSetMap (fold_left f' (a :: l) m) (f' (fold_left f' l m) a)). intro. +generalize (H1 (fst_ext e')). clear H1. intro H1. inversion H1; clear H1. +rewrite <-H3 in *. elim (VertexSet.empty_1 H0). +set (tmp := fold_left f' l m) in *. +unfold f' in H3. unfold f in H3. +destruct (Vertex.eq_dec (fst_ext e') a). +rewrite MapFacts.add_eq_o in H3. inversion H3; subst; clear H3. +rewrite <-H2 in H0. +rewrite H4 in H0. clear H4. +destruct (proj1 (Props.Dec.F.add_iff _ _ _ ) H0). +exists (fst_ext e', snd_ext e, Some N0). +split. +left. apply (proj2 (edgemap_to_edgeset_charac _ _ _ (Some N0) (sym_pmap g))). +apply (sym_pmap g). +assert (VertexSet.In (fst_ext e') s'). +apply HH. left. auto. +unfold s' in H3. apply (VertexSet.remove_3 H3). +unfold redirect. change_rewrite. +destruct (OTFacts.eq_dec (fst_ext e') (snd_ext e)). +unfold weak_eq. change_rewrite. +split. +right. split; auto. +unfold same_type. left. split; exists N0; auto. +destruct (OTFacts.eq_dec (snd_ext e) (snd_ext e)). +unfold weak_eq. change_rewrite. +split. +left. split; auto. +left. split; exists N0; auto. +elim n2. intuition. +exists (fst_ext e', snd_ext e', Some N0). +split. +left. +apply (proj2 (edgemap_to_edgeset_charac _ _ _ (Some N0) (sym_pmap g))). +generalize (VertexSet.remove_3 H1). clear H1. intro H1. +unfold adj_set. rewrite (MapFacts.find_o _ e0). assumption. +unfold redirect. change_rewrite. +destruct (OTFacts.eq_dec (fst_ext e') (snd_ext e)). +elim (merge_1 e g p q). +rewrite <-r. apply (proj1 (In_graph_edge_in_ext _ _ Hin)). +destruct (OTFacts.eq_dec (snd_ext e') (snd_ext e)). +elim (merge_1 e g p q). +rewrite <-r. apply (proj2 (In_graph_edge_in_ext _ _ Hin)). +unfold weak_eq. change_rewrite. +split. +left. split; auto. +left. split; exists N0; auto. +intuition. +auto. + +rewrite MapFacts.add_neq_o in H3. rewrite <-H3 in IHl. +apply IHl. rewrite <-H2 in H0. rewrite H4 in H0. auto. +intros. apply HH. auto. +auto. + +apply fold_left_assoc_map. +unfold EqualSetMap. unfold f'. unfold f. +intros. +destruct (Vertex.eq_dec x z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x y). +rewrite MapFacts.add_eq_o. +constructor. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). +rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +apply VertexSet.eq_refl. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x y). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x s0); constructor; apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f'. unfold f. +destruct (Vertex.eq_dec x a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H1. auto. auto. + +auto. +auto. + +generalize (IE_weights _ _ H0). intro Hw. +unfold IE in H0. simpl in H0. +rewrite (edge_eq e') in H0. rewrite Hw in H0. +generalize (proj1 (edgemap_to_edgeset_charac _ _ _ None (sym_imap_merge_map _ _ q p)) H0). +clear H0. intro H0. +unfold imap_merge in H0. unfold map_merge in H0. +set (s := (VertexSet.union + (VertexSet.remove (snd_ext e) (adj_set (fst_ext e) (imap g))) + (VertexSet.remove (fst_ext e) (adj_set (snd_ext e) (imap g))))) in *. +set (m := imap g) in *. +set (f := (fun (y : VertexSet.elt) (m' : VertexMap.t VertexSet.t) => + VertexMap.add y + (VertexSet.add (fst_ext e) + (VertexSet.remove (snd_ext e) (adj_set y m))) m')) in *. +set (s' := VertexSet.remove (fst_ext e) (adj_set (snd_ext e) m)) in *. +unfold adj_set in H0. +destruct (Vertex.eq_dec (fst_ext e') (snd_ext e)). +rewrite MapFacts.remove_eq_o in H0. +elim (VertexSet.empty_1 H0). intuition. +rewrite MapFacts.remove_neq_o in H0. +destruct (Vertex.eq_dec (fst_ext e') (fst_ext e)). +rewrite MapFacts.add_eq_o in H0. +unfold s in H0. +destruct (VertexSet.union_1 H0). +generalize (VertexSet.remove_3 H1). clear H1. intro H1. +exists (fst_ext e', snd_ext e', None). split. +right. unfold IE. +apply (proj2 (edgemap_to_edgeset_charac _ _ _ None (sym_imap g))). +unfold adj_set. rewrite (MapFacts.find_o _ e0). assumption. +unfold redirect. change_rewrite. +destruct (OTFacts.eq_dec (fst_ext e') (snd_ext e)). +elim (n r). +destruct (OTFacts.eq_dec (snd_ext e') (snd_ext e)). +elim (merge_1 e g p q). +rewrite <-r. apply (proj2 (In_graph_edge_in_ext _ _ Hin)). +unfold weak_eq. change_rewrite. +split. left; split; auto. +right; split; unfold interf_edge; simpl; auto. +unfold s' in H1. +generalize (VertexSet.remove_3 H1). intro H2. +exists (snd_ext e', snd_ext e, None). +split. +right. unfold IE. +apply (proj2 (edgemap_to_edgeset_charac _ _ _ None (sym_imap g))). +apply (sym_imap g). assumption. +unfold redirect. change_rewrite. +destruct (OTFacts.eq_dec (snd_ext e') (snd_ext e)). +unfold weak_eq. change_rewrite. +split. +left. split; auto. +right. unfold interf_edge;split; simpl; auto. +destruct (OTFacts.eq_dec (snd_ext e) (snd_ext e)). +unfold weak_eq. change_rewrite. +split. +right. split; auto. +right. unfold interf_edge; split; simpl; auto. +elim n1. intuition. +intuition. +rewrite MapFacts.add_neq_o in H0. + +rewrite VertexSet.fold_1 in H0. +set (f' := fun a e => f e a) in *. +generalize VertexSet.elements_2. intro HH. +generalize (HH s'). clear HH. intro HH. +induction (VertexSet.elements s'). simpl in H0. +case_eq (VertexMap.find (fst_ext e') m); intros. rewrite H1 in H0. +exists (fst_ext e', snd_ext e', None). +split. +right. apply (proj2 (edgemap_to_edgeset_charac _ _ _ None (sym_imap g))). +unfold adj_set. fold m. rewrite H1. assumption. +unfold redirect. change_rewrite. +destruct (OTFacts.eq_dec (fst_ext e') (snd_ext e)). +elim (n r). +destruct (OTFacts.eq_dec (snd_ext e') (snd_ext e)). +elim (merge_1 e g p q). +rewrite <-r. apply (proj2 (In_graph_edge_in_ext _ _ Hin)). +unfold weak_eq. change_rewrite. +split. +left. split; auto. +right. unfold interf_edge; split; simpl; auto. +rewrite H1 in H0. elim (VertexSet.empty_1 H0). +simpl in H0. +cut (EqualSetMap (fold_left f' (a :: l) m) (f' (fold_left f' l m) a)). intro. +generalize (H1 (fst_ext e')). clear H1. intro H1. inversion H1; clear H1. +rewrite <-H3 in *. elim (VertexSet.empty_1 H0). +set (tmp := fold_left f' l m) in *. +unfold f' in H3. unfold f in H3. +destruct (Vertex.eq_dec (fst_ext e') a). +rewrite MapFacts.add_eq_o in H3. inversion H3; subst; clear H3. +rewrite <-H2 in H0. +rewrite H4 in H0. clear H4. +destruct (proj1 (Props.Dec.F.add_iff _ _ _ ) H0). +exists (fst_ext e', snd_ext e, None). +split. +right. apply (proj2 (edgemap_to_edgeset_charac _ _ _ None (sym_imap g))). +apply (sym_imap g). +assert (VertexSet.In (fst_ext e') s'). +apply HH. left. auto. +unfold s' in H3. apply (VertexSet.remove_3 H3). +unfold redirect. change_rewrite. +destruct (OTFacts.eq_dec (fst_ext e') (snd_ext e)). +unfold weak_eq. change_rewrite. +split. +right. split; auto. +right. unfold interf_edge; split; simpl; auto. +destruct (OTFacts.eq_dec (snd_ext e) (snd_ext e)). +unfold weak_eq. change_rewrite. +split. +left. split; auto. +right. unfold interf_edge; split; simpl; auto. +elim n2. intuition. +exists (fst_ext e', snd_ext e', None). +split. +right. apply (proj2 (edgemap_to_edgeset_charac _ _ _ None (sym_imap g))). +generalize (VertexSet.remove_3 H1). clear H1. intro H1. +unfold adj_set. rewrite (MapFacts.find_o _ e0). assumption. +unfold redirect. change_rewrite. +destruct (OTFacts.eq_dec (fst_ext e') (snd_ext e)). +elim (merge_1 e g p q). +rewrite <-r. apply (proj1 (In_graph_edge_in_ext _ _ Hin)). +destruct (OTFacts.eq_dec (snd_ext e') (snd_ext e)). +elim (merge_1 e g p q). +rewrite <-r. apply (proj2 (In_graph_edge_in_ext _ _ Hin)). +unfold weak_eq. change_rewrite. +split. +left. split; auto. +right. unfold interf_edge; split; simpl; auto. +intuition. + +rewrite MapFacts.add_neq_o in H3. rewrite <-H3 in IHl. +apply IHl. rewrite <-H2 in H0. rewrite H4 in H0. auto. +intros. apply HH. auto. +auto. + +apply fold_left_assoc_map. +unfold EqualSetMap. unfold f'. unfold f. +intros. +destruct (Vertex.eq_dec x z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x y). +rewrite MapFacts.add_eq_o. +constructor. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). +rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +apply VertexSet.eq_refl. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x y). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x s0); constructor; apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f'. unfold f. +destruct (Vertex.eq_dec x a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H1. auto. auto. + +auto. +auto. +Qed. + +Lemma resolve_conflicts_map_5 : forall x y e g, +aff_edge e -> +In_graph_edge e g -> +VertexSet.In x (adj_set y (map_merge e (pmap g))) -> +~VertexSet.In x (adj_set y (imap_merge e g)) -> +VertexSet.In x (adj_set y (resolve_conflicts (fst_ext e) (map_merge e (pmap g)) + (adj_set (fst_ext e) (map_merge e (pmap g))) + (adj_set (fst_ext e) (imap_merge e g)))). + +Proof. +intros x y e g p q H H0. +unfold resolve_conflicts. +set (f := (fun (x0 : VertexSet.elt) (m : VertexMap.t VertexSet.t) => + VertexMap.add x0 + (VertexSet.remove (fst_ext e) + (adj_set x0 (map_merge e (pmap g)))) m)) in *. +set (m := map_merge e (pmap g)) in *. +set (s' := (VertexSet.diff (adj_set (fst_ext e) m) + (adj_set (fst_ext e) (imap_merge e g)))) in *. +set (s := (VertexSet.inter (adj_set (fst_ext e) m) + (adj_set (fst_ext e) (imap_merge e g)))) in *. +rewrite VertexSet.fold_1. +set (f' := fun a e => f e a) in *. +generalize VertexSet.elements_1. intro. +generalize (H1 s y). clear H1. intro HH. +generalize VertexSet.elements_2. intro. +generalize (H1 s y). clear H1. intro HHH. +induction (VertexSet.elements s). simpl. +destruct (Vertex.eq_dec y (fst_ext e)). +unfold adj_set. rewrite (MapFacts.find_o _ e0). +rewrite MapFacts.add_eq_o. +apply VertexSet.diff_3. +unfold adj_set. rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). assumption. +unfold adj_set. rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). assumption. +intuition. +unfold adj_set. rewrite MapFacts.add_neq_o. assumption. +auto. + +cut (EqualSetMap (fold_left f' (a :: l) m) (f' (fold_left f' l m) a)). +intro. generalize (H1 y). clear H1. intro H1. simpl. inversion H1; clear H1. +set (tmp := fold_left f' l m) in *. +unfold f' in H4. unfold f in H4. +destruct (Vertex.eq_dec y a). +rewrite MapFacts.add_eq_o in H4. congruence. intuition. +destruct (Vertex.eq_dec y (fst_ext e)). unfold adj_set. +rewrite MapFacts.add_eq_o. +unfold adj_set in IHl. +rewrite MapFacts.add_eq_o in IHl. +apply IHl. intro. +generalize (HH H1). intro. +inversion H2; subst. +elim (n H6). +auto. +intro. apply HHH. right. auto. +intuition. +intuition. + +rewrite MapFacts.add_neq_o in H4. +unfold adj_set in IHl. rewrite MapFacts.add_neq_o in IHl. +rewrite <-H4 in IHl. +assert (VertexSet.In x VertexSet.empty). +apply IHl. intro. +generalize (HH H1). intro. +inversion H2; subst. +elim (n H6). +auto. +intro. apply HHH. right. auto. +elim (VertexSet.empty_1 H1). +auto. +auto. +auto. + +set (tmp := fold_left f' l m) in *. +unfold f' in H3. unfold f in H3. +destruct (Vertex.eq_dec y a). +rewrite MapFacts.add_eq_o in H3. +unfold adj_set. destruct (Vertex.eq_dec y (fst_ext e)). +rewrite MapFacts.add_eq_o. +apply VertexSet.diff_3. +unfold adj_set. rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). assumption. +unfold adj_set. rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). assumption. +intuition. + +rewrite MapFacts.add_neq_o. rewrite <-H2. rewrite H4. +inversion H3; subst; clear H3. +apply VertexSet.remove_2. + +intro H5. +assert (VertexSet.In y s). +apply HHH. left. auto. +generalize (VertexSet.inter_2 H1). intro. +unfold adj_set in H3. rewrite (MapFacts.find_o _ H5) in H3. +generalize (sym_imap_merge_map e g p q _ _ H3). intro. +elim H0. assumption. +unfold adj_set. rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +auto. +auto. +intuition. + +rewrite MapFacts.add_neq_o in H3. +destruct (Vertex.eq_dec y (fst_ext e)). +unfold adj_set in IHl. +rewrite MapFacts.add_eq_o in IHl. +unfold adj_set. rewrite MapFacts.add_eq_o. +apply IHl. + +intro. generalize (HH H1). intro. +inversion H5; subst. +elim (n H7). +auto. + +intro. apply HHH. right. auto. +intuition. +intuition. + +unfold adj_set in IHl. rewrite MapFacts.add_neq_o in IHl. +unfold adj_set. rewrite MapFacts.add_neq_o. +rewrite <-H2. +rewrite <-H3 in IHl. +rewrite H4. +apply IHl. + +intro. generalize (HH H1). intro. +inversion H5; subst. +elim (n H7). +auto. + +intro. apply HHH. right. auto. +auto. +auto. +auto. + +apply fold_left_assoc_map. +unfold EqualSetMap. unfold f'. unfold f. +intros. +destruct (Vertex.eq_dec x0 z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +constructor. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). +rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +apply VertexSet.eq_refl. intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x0 y0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x0 s0); constructor; apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f'. unfold f. +destruct (Vertex.eq_dec x0 a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H1. auto. auto. +Qed. + +Lemma merge_3_aux : forall e e' g, +aff_edge e -> +In_graph_edge e g -> +In_graph_edge e' g -> +aff_edge e' -> +~Edge.eq e e' -> +VertexSet.In (snd_ext (redirect (snd_ext e) (fst_ext e) e')) + (adj_set (fst_ext (redirect (snd_ext e) (fst_ext e) e')) (map_merge e (pmap g))). + +Proof. +intros e e' g q p H0 H1 Hdiff. generalize I. intro H. +assert (get_weight e = Some N0) as Hw. +destruct p. apply (AE_weights _ _ H2). +generalize (IE_weights _ _ H2). inversion q. congruence. +assert (get_weight e' = Some N0) as Hw'. +destruct H0. apply (AE_weights _ _ H0). +generalize (IE_weights _ _ H0). inversion H1. congruence. +assert (EdgeSet.In (fst_ext e, snd_ext e, Some N0) (AE g)) as He. +destruct p. rewrite (edge_eq e) in H2. +generalize (AE_weights _ _ H2). intro. simpl in H3. +rewrite H3 in H2. assumption. +generalize (IE_weights _ _ H2). inversion q. congruence. +assert (VertexSet.In (snd_ext e) (adj_set (fst_ext e) (pmap g))) as Hee. +apply (proj1 (edgemap_to_edgeset_charac _ _ _ (Some N0) (sym_pmap g))). +assumption. +assert (EdgeSet.In (fst_ext e', snd_ext e', Some N0) (AE g)) as He'. +destruct H0. rewrite (edge_eq e') in H0. +generalize (AE_weights _ _ H0). intro. simpl in H2. +rewrite H2 in H0. assumption. +generalize (IE_weights _ _ H0). intro H2. inversion H1. congruence. +assert (VertexSet.In (snd_ext e') (adj_set (fst_ext e') (pmap g))) as Hee'. +apply (proj1 (edgemap_to_edgeset_charac _ _ _ (Some N0) (sym_pmap g))). +assumption. +assert (forall a b, VertexSet.In a (adj_set b (map_merge e (pmap g))) -> + VertexSet.In b (adj_set a (map_merge e (pmap g)))) as Hsym. +intros; apply sym_map_merge_pmap; auto. + +unfold redirect. +destruct (OTFacts.eq_dec (fst_ext e') (snd_ext e)). +unfold map_merge. +set (s := (VertexSet.union + (VertexSet.remove (snd_ext e) (adj_set (fst_ext e) (pmap g))) + (VertexSet.remove (fst_ext e) (adj_set (snd_ext e) (pmap g))))) in *. +set (m := pmap g) in *. +set (f := (fun (y : VertexSet.elt) (m' : VertexMap.t VertexSet.t) => + VertexMap.add y + (VertexSet.add (fst_ext e) + (VertexSet.remove (snd_ext e) (adj_set y m))) m')) in *. +set (s' := VertexSet.remove (fst_ext e) (adj_set (snd_ext e) m)) in *. +unfold adj_set. +rewrite MapFacts.remove_neq_o. +rewrite MapFacts.add_eq_o. +unfold s. +apply VertexSet.union_3. +unfold s'. +apply VertexSet.remove_2. +intro. change_rewrite. +elim Hdiff. +rewrite (edge_eq e). rewrite edge_comm. apply eq_ordered_eq. +rewrite (edge_eq e'). unfold E.eq; change_rewrite; simpl; intuition. +rewrite Hw. rewrite Hw'. apply OptionN_as_OT.eq_refl. +unfold adj_set. rewrite (MapFacts.find_o _ (Vertex.eq_sym r)). +fold (adj_set (fst_ext e') m). +assumption. change_rewrite. intuition. +change_rewrite. +apply (In_graph_edge_diff_ext _ _ p). + +destruct (OTFacts.eq_dec (snd_ext e') (snd_ext e)); change_rewrite. +unfold map_merge. +set (s := (VertexSet.union + (VertexSet.remove (snd_ext e) (adj_set (fst_ext e) (pmap g))) + (VertexSet.remove (fst_ext e) (adj_set (snd_ext e) (pmap g))))) in *. +set (m := pmap g) in *. +set (f := (fun (y : VertexSet.elt) (m' : VertexMap.t VertexSet.t) => + VertexMap.add y + (VertexSet.add (fst_ext e) + (VertexSet.remove (snd_ext e) (adj_set y m))) m')) in *. +set (s' := VertexSet.remove (fst_ext e) (adj_set (snd_ext e) m)) in *. +unfold adj_set. +rewrite MapFacts.remove_neq_o. +rewrite MapFacts.add_neq_o. +rewrite VertexSet.fold_1. +set (f' := fun a e => f e a). +generalize VertexSet.elements_1. intro HH. +generalize (HH s' (fst_ext e')). clear HH. intro HH. +induction (VertexSet.elements s'). simpl. +assert (InA Vertex.eq (fst_ext e') nil). +apply HH. +unfold s'. apply VertexSet.remove_2. +intro. elim Hdiff. +apply eq_ordered_eq. +rewrite (edge_eq e). rewrite (edge_eq e'). +unfold E.eq; change_rewrite; simpl; intuition. +rewrite Hw. rewrite Hw'. apply OptionN_as_OT.eq_refl. +unfold adj_set. rewrite (MapFacts.find_o _ (Vertex.eq_sym r)). +apply (sym_pmap g). assumption. +inversion H2. + +cut (EqualSetMap (fold_left f' (a :: l) m) (f' (fold_left f' l m) a)). intro. +generalize (H2 (fst_ext e')). clear H2. intro H2. simpl. inversion H2;clear H2. +set (tmp := fold_left f' l m) in *. +unfold f' in H5. unfold f in H5. +destruct (Vertex.eq_dec (fst_ext e') a). +rewrite MapFacts.add_eq_o in H5. congruence. intuition. +rewrite MapFacts.add_neq_o in H5. rewrite <-H5 in IHl. +apply IHl. +intro. generalize (HH H2). clear HH. intro HH. +inversion HH; subst. +elim (n0 H6). +assumption. +auto. + +set (tmp := fold_left f' l m) in *. +unfold f' in H4. unfold f in H4. +destruct (Vertex.eq_dec (fst_ext e') a). +rewrite MapFacts.add_eq_o in H4. +rewrite H5. inversion H4; subst; clear H4. +apply VertexSet.add_1. intuition. intuition. +rewrite MapFacts.add_neq_o in H4. +rewrite <-H4 in IHl. +rewrite H5. apply IHl. +intro. generalize (HH H2). clear HH. intro HH. inversion HH; subst. +elim (n0 H7). +assumption. +auto. + +apply fold_left_assoc_map. +unfold EqualSetMap. unfold f'. unfold f. +intros. +destruct (Vertex.eq_dec x z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x y). +rewrite MapFacts.add_eq_o. +constructor. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). +rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +apply VertexSet.eq_refl. intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x y). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x s0); constructor; apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f'. unfold f. +destruct (Vertex.eq_dec x a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H2. auto. auto. + +intro. elim Hdiff. +apply eq_ordered_eq. +rewrite (edge_eq e). rewrite (edge_eq e'). +unfold E.eq; change_rewrite; simpl; intuition. +rewrite Hw. rewrite Hw'. apply OptionN_as_OT.eq_refl. + +intuition. +unfold map_merge. +set (s := (VertexSet.union + (VertexSet.remove (snd_ext e) (adj_set (fst_ext e) (pmap g))) + (VertexSet.remove (fst_ext e) (adj_set (snd_ext e) (pmap g))))) in *. +set (m := pmap g) in *. +set (f := (fun (y : VertexSet.elt) (m' : VertexMap.t VertexSet.t) => + VertexMap.add y + (VertexSet.add (fst_ext e) + (VertexSet.remove (snd_ext e) (adj_set y m))) m')) in *. +set (s' := VertexSet.remove (fst_ext e) (adj_set (snd_ext e) m)) in *. +unfold adj_set. +rewrite MapFacts.remove_neq_o. +destruct (Vertex.eq_dec (fst_ext e') (fst_ext e)). +rewrite MapFacts.add_eq_o. +unfold s. +apply VertexSet.union_2. +apply VertexSet.remove_2. intuition. +unfold adj_set. rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +assumption. +intuition. + +rewrite MapFacts.add_neq_o. +rewrite VertexSet.fold_1. +set (f' := fun a e => f e a). +(* +generalize VertexSet.elements_1. intro HH. +generalize (HH s' (fst_ext e')). clear HH. intro HH. +*) +induction (VertexSet.elements s'). simpl. assumption. +cut (EqualSetMap (fold_left f' (a :: l) m) (f' (fold_left f' l m) a)). intro. +generalize (H2 (fst_ext e')). clear H2. intro H2. simpl. inversion H2;clear H2. +set (tmp := fold_left f' l m) in *. +unfold f' in H5. unfold f in H5. +destruct (Vertex.eq_dec (fst_ext e') a). +rewrite MapFacts.add_eq_o in H5. congruence. intuition. +rewrite MapFacts.add_neq_o in H5. rewrite <-H5 in IHl. +apply IHl. auto. + +set (tmp := fold_left f' l m) in *. +unfold f' in H4. unfold f in H4. +destruct (Vertex.eq_dec (fst_ext e') a). +rewrite MapFacts.add_eq_o in H4. +rewrite H5. inversion H4; subst; clear H4. +apply VertexSet.add_2. apply VertexSet.remove_2. intuition. +unfold adj_set. rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). assumption. +intuition. + +rewrite MapFacts.add_neq_o in H4. +rewrite <-H4 in IHl. +rewrite H5. apply IHl. +auto. + +apply fold_left_assoc_map. +unfold EqualSetMap. unfold f'. unfold f. +intros. +destruct (Vertex.eq_dec x z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x y). +rewrite MapFacts.add_eq_o. +constructor. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_sym e1)). +rewrite (MapFacts.find_o _ (Vertex.eq_sym e0)). +apply VertexSet.eq_refl. intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x y). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x s0); constructor; apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +unfold f'. unfold f. +destruct (Vertex.eq_dec x a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H2. auto. auto. + +auto. +intuition. +Qed. + +Lemma merge_5 : forall e e' g p q, +In_graph_edge e' g -> +~Edge.eq e e' -> +aff_edge e' -> +~Interfere (fst_ext (redirect (snd_ext e) (fst_ext e) e')) + (snd_ext (redirect (snd_ext e) (fst_ext e) e')) + (merge e g p q) -> +Prefere (fst_ext (redirect (snd_ext e) (fst_ext e) e')) + (snd_ext (redirect (snd_ext e) (fst_ext e) e')) + (merge e g p q). + +Proof. +intros e e' g p q H0 H1 H2 H3. generalize I. intro H. +unfold Prefere. exists N0. left. +unfold merge. unfold AE. simpl. +apply (proj2 (edgemap_to_edgeset_charac _ _ _ (Some N0) (sym_pmap_merge_map e g q p))). +apply resolve_conflicts_map_5. +assumption. +assumption. +apply merge_3_aux; assumption. +intro. elim H3. unfold Interfere. right. unfold merge. +unfold IE. simpl. +apply (proj2 (edgemap_to_edgeset_charac _ _ _ None (sym_imap_merge_map e g q p))). +assumption. +Qed. +(* +unfold change. +destruct (OTFacts.eq_dec (fst_ext e') (snd_ext e)). change_rewrite. +apply VertexSet.remove_2. +apply (In_graph_edge_diff_ext e g p). +apply (proj1 (In_graph_edge_in_ext _ _ p)). +destruct (OTFacts.eq_dec (snd_ext e') (snd_ext e)). change_rewrite. +apply VertexSet.remove_2. +intro. elim (In_graph_edge_diff_ext e' g H0). +apply Vertex.eq_trans with (y := snd_ext e); auto. +apply (proj1 (In_graph_edge_in_ext _ _ H0)). +apply VertexSet.remove_2. +auto. +apply (proj1 (In_graph_edge_in_ext _ _ H0)). +Qed. +*) + +Lemma precolored_merge2 : forall x e g p q, +(VertexSet.In x (VertexSet.remove (snd_ext e) (precolored g)) <-> + VertexSet.In x (precolored (merge e g p q))). + +Proof. +intros x e g HH q. split; intros. +unfold merge. +unfold precolored. simpl. unfold is_precolored. simpl. +apply VertexSet.filter_3. +unfold compat_bool;intros. +rewrite (mreg_ext _ _ H0). auto. +apply VertexSet.remove_2. +intro H1. elim (VertexSet.remove_1 H1 H). +apply (VertexSet.filter_1 (compat_bool_is_precolored _) (VertexSet.remove_3 H)). +generalize (VertexSet.filter_2 (compat_bool_is_precolored _) (VertexSet.remove_3 H)). intro H0. +assumption. + +unfold precolored, is_precolored in *. unfold merge in H. simpl in H. +apply VertexSet.remove_2. intro. +generalize (VertexSet.filter_1 (compat_bool_is_precolored (merge e g HH q)) H). intro. +elim (VertexSet.remove_1 H0 H1). +apply VertexSet.filter_3. +unfold compat_bool. apply mreg_ext. +generalize (VertexSet.filter_1 (compat_bool_is_precolored (merge e g HH q)) H). intros. +apply (VertexSet.remove_3 H0). +apply (VertexSet.filter_2 (compat_bool_is_precolored (merge e g HH q)) H). +Qed. + +(* spec of delete_preference_edges *) +Lemma delete_preference_edges_prec : forall r g Hin, +VertexSet.Equal (precolored (delete_preference_edges r g Hin)) (precolored g). + +Proof. +intros. +unfold delete_preference_edges. +unfold precolored. simpl. +unfold is_precolored. simpl. +apply VertexSet.eq_refl. +Qed. + +Lemma delete_preference_edges_1 : forall e r g Hin, +In_graph_edge e (delete_preference_edges r g Hin) -> In_graph_edge e g. + +Proof. +intros. +unfold delete_preference_edges in H. +destruct H;[left; generalize (AE_weights _ _ H); unfold AE in *| + right; generalize (IE_weights _ _ H); unfold IE in *]; simpl in H; intros. +rewrite (edge_eq e) in H. +rewrite H0 in H. generalize (proj1 (edgemap_to_edgeset_charac _ _ _ _(sym_pmap_delete_preferences r g)) H). intro. +rewrite (edge_eq e). +rewrite H0. apply (proj2 (edgemap_to_edgeset_charac _ _ _ _ (sym_pmap g))). +apply (delete_preference_sub _ _ _ _ H1). +assumption. +Qed. + +Lemma IE_delete_preference_eq : forall x g H, +EdgeSet.Equal (IE g) (IE (delete_preference_edges x g H)). + +Proof. +intros. +unfold delete_preference_edges. +destruct (In_graph_dec x g). +unfold IE. simpl. apply EdgeSet.eq_refl. +apply EdgeSet.eq_refl. +Qed. + +Lemma delete_preference_edges_2 : forall e r g Hin, +In_graph_edge e g -> +~incident e r -> +In_graph_edge e (delete_preference_edges r g Hin). + +Proof. +intros. +destruct H;[left|right]. +generalize (AE_weights _ _ H). intro Hw. +unfold AE in *. +rewrite (edge_eq e). rewrite Hw. +unfold delete_preference_edges. +apply (proj2 (edgemap_to_edgeset_charac _ _ _ _ (sym_pmap_delete_preferences r g))). +unfold pmap_delete_preferences. +rewrite VertexSet.fold_1. +induction (VertexSet.elements (adj_set r (pmap g))). simpl. +unfold adj_set. rewrite InterfFacts.add_neq_o. +rewrite (edge_eq e) in H. rewrite Hw in H. +apply (proj1 (edgemap_to_edgeset_charac _ _ _ _ (sym_pmap g)) H). +intro H1. elim H0. left. auto. + +cut (EqualSetMap + (fold_left + (fun (a : VertexMap.t VertexSet.t) (e : VertexSet.elt) => + VertexMap.add e (VertexSet.remove r (adj_set e (pmap g))) a) (a :: l) + (VertexMap.add r VertexSet.empty (pmap g))) + (VertexMap.add a + (VertexSet.remove r + (adj_set a (pmap g))) + (fold_left + (fun (a : VertexMap.t VertexSet.t) (e : VertexSet.elt) => + VertexMap.add e (VertexSet.remove r (adj_set e (pmap g))) a) l + (VertexMap.add r VertexSet.empty (pmap g))))). intro. +unfold EqualSetMap in H1. generalize (H1 (fst_ext e)). clear H1. intro H1. +inversion H1. simpl. unfold adj_set in *. rewrite <-H3. +destruct (Vertex.eq_dec (fst_ext e) a). +rewrite InterfFacts.add_eq_o in H4. inversion H4. +intuition. +rewrite InterfFacts.add_neq_o in H4. rewrite <-H4 in *. +elim (VertexSet.empty_1 IHl). +auto. +simpl. unfold adj_set in *. rewrite <-H2. +rewrite H4. clear H1. clear H2. clear H4. +destruct (Vertex.eq_dec (fst_ext e) a). +rewrite InterfFacts.add_eq_o in H3. +inversion H3. clear H2. +apply VertexSet.remove_2. +intro. elim H0. right. auto. +clear H3. +rewrite (edge_eq e) in H. rewrite Hw in H. +rewrite (InterfFacts.find_o _ (Vertex.eq_sym e0)). +apply (proj1 (edgemap_to_edgeset_charac _ _ _ _(sym_pmap g)) H). +intuition. +rewrite (InterfFacts.add_neq_o) in H3. +case_eq (VertexMap.find (elt:=VertexSet.t) (fst_ext e) + (fold_left + (fun (a : VertexMap.t VertexSet.t) (e : VertexSet.elt) => + VertexMap.add e + (VertexSet.remove r + match VertexMap.find (elt:=VertexSet.t) e (pmap g) with + | Some x => x + | None => VertexSet.empty + end) a) l (VertexMap.add r VertexSet.empty (pmap g)))); +intros; rewrite H1 in *. +inversion H3. assumption. +inversion H3. +auto. + +apply fold_left_assoc_map. +intros. +unfold EqualSetMap. intros. +destruct (Vertex.eq_dec x z). +rewrite InterfFacts.add_eq_o. +destruct (Vertex.eq_dec x y). +rewrite InterfFacts.add_eq_o. +constructor. +generalize (Vertex.eq_trans (Vertex.eq_sym e0) e1). intro HHH. +unfold adj_set. +rewrite (InterfFacts.find_o _ HHH). apply VertexSet.eq_refl. +intuition. + +rewrite InterfFacts.add_neq_o. +rewrite InterfFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +auto. +intuition. + +rewrite InterfFacts.add_neq_o. +destruct (Vertex.eq_dec x y). +rewrite InterfFacts.add_eq_o. +rewrite InterfFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. + +intros. +unfold EqualSetMap. intros. +destruct (Vertex.eq_dec x y). +rewrite InterfFacts.add_eq_o. +rewrite InterfFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. +rewrite InterfFacts.add_neq_o. +rewrite InterfFacts.add_neq_o. +rewrite InterfFacts.add_neq_o. +apply EqualSetMap_refl. +auto. +auto. +auto. +auto. + +unfold EqualSetMap. intros. +destruct (Vertex.eq_dec x a0). +rewrite InterfFacts.add_eq_o. +rewrite InterfFacts.add_eq_o. +constructor. apply VertexSet.eq_refl. +intuition. +intuition. +rewrite InterfFacts.add_neq_o. +rewrite InterfFacts.add_neq_o. +apply (H1 x). +auto. +auto. + +rewrite <-IE_delete_preference_eq. auto. +Qed. + +Lemma V_delete_preference_eq : forall x g Hin, +VertexSet.Equal (V g) (V (delete_preference_edges x g Hin)). + +Proof. +intros. +unfold delete_preference_edges. +destruct (In_graph_dec x g); simpl; apply VertexSet.eq_refl. +Qed. + +Lemma in_delete_preference_not_incident : forall e r g Hdep, +EdgeSet.In e (AE (delete_preference_edges r g Hdep)) -> ~incident e r. + +Proof. +intros. +generalize (AE_weights _ _ H). intro Hw. +generalize (proj1 (In_graph_aff_edge_in_AE _ _) H). intro Hin. +unfold AE in H. +rewrite (edge_eq e) in H. rewrite Hw in H. +unfold delete_preference_edges in H. +generalize (proj1 (edgemap_to_edgeset_charac _ _ _ _(sym_pmap_delete_preferences r g)) H). clear H. intro H. +unfold pmap_delete_preferences in H. +rewrite VertexSet.fold_1 in H. +generalize VertexSet.elements_1. intro HH. +generalize (HH (adj_set r (pmap g)) (fst_ext e)). clear HH. intro HH. +generalize VertexSet.elements_2. intro HH1. +generalize (HH1 (adj_set r (pmap g))). clear HH1. intro HH1. + +assert (forall z, In z (VertexSet.elements (adj_set r (pmap g))) -> ~Vertex.eq r z) as Hneq. +clear H HH Hin Hw Hdep. +intros. intro H0. +induction (VertexSet.elements (adj_set r (pmap g))). inversion H. +simpl in H. destruct H. subst. +assert (VertexSet.In r (adj_set r (pmap g))). +apply HH1. left. auto. +elim (not_eq_extremities g r r). +right. assumption. +auto. +apply IHl. +intros. apply HH1. auto. +assumption. + +induction (VertexSet.elements (adj_set r (pmap g))). simpl in H. +unfold adj_set in H. +destruct (Vertex.eq_dec (fst_ext e) r). +rewrite InterfFacts.add_eq_o in H. +elim (VertexSet.empty_1 H). intuition. +rewrite InterfFacts.add_neq_o in H. +destruct (Vertex.eq_dec (snd_ext e) r). +assert (InA Vertex.eq (fst_ext e) nil). +apply HH. +apply (proj1 (edgemap_to_edgeset_charac (pmap g) r (fst_ext e) (Some N0) (sym_pmap g))). +assert (eq (r,fst_ext e, Some N0) e). +rewrite edge_comm. +apply eq_ordered_eq. +unfold E.eq; simpl; intuition. apply Regs.eq_refl. +rewrite <-Hw. apply OptionN_as_OT.eq_refl. +rewrite H0. +destruct Hin. +generalize (delete_preference_edges_1 e r g Hdep H2). clear H2. intro H2. +destruct H2. +unfold AE in H2. assumption. +generalize (IE_weights _ _ H2). intro Hw'. rewrite Hw' in Hw. inversion Hw. +inversion H0. +intro H0. destruct H0. +elim n. auto. +elim n0. auto. +auto. + +set (f := (fun (a : VertexMap.t VertexSet.t) (e : VertexSet.elt) => + VertexMap.add e (VertexSet.remove r (adj_set e (pmap g))) a)) in *. +set (s := VertexMap.add r VertexSet.empty (pmap g)) in *. + +assert (EqualSetMap (fold_left f (a :: l) s) (f (fold_left f l s) a)). +apply fold_left_assoc_map. + +unfold f. unfold EqualSetMap. intros. +destruct (Vertex.eq_dec x z). +rewrite MapFacts.add_eq_o. +destruct (Vertex.eq_dec x y). +rewrite MapFacts.add_eq_o. +constructor. +apply Props.Equal_remove. +unfold adj_set. +rewrite (MapFacts.find_o _ (Vertex.eq_trans (Vertex.eq_sym e0) e1)). +apply VertexSet.eq_refl. intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +auto. +intuition. +rewrite MapFacts.add_neq_o. +destruct (Vertex.eq_dec x y). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +destruct (VertexMap.find x s0); constructor. +apply VertexSet.eq_refl. +auto. +auto. +auto. +auto. + +unfold f. unfold EqualSetMap. intros. +destruct (Vertex.eq_dec x a0). +rewrite MapFacts.add_eq_o. +rewrite MapFacts.add_eq_o. +constructor. +apply VertexSet.eq_refl. +intuition. +intuition. +rewrite MapFacts.add_neq_o. +rewrite MapFacts.add_neq_o. +apply H0. +auto. +auto. + +generalize (H0 (fst_ext e)). clear H0. intro H0. +simpl in H. unfold adj_set in H. +inversion H0. subst. +rewrite <-H2 in H. +elim (VertexSet.empty_1 H). +rewrite <-H1 in H. +set (tmp := fold_left f l s) in *. +unfold f in H2. +destruct (Vertex.eq_dec (fst_ext e) a). +rewrite MapFacts.add_eq_o in H2. +intro. +inversion H2. subst. clear H2. +rewrite H3 in H. +destruct H4. + +elim (Hneq a). +left. auto. +apply (Vertex.eq_trans H2 e0). +elim (VertexSet.remove_1 H2 H). +intuition. + +rewrite MapFacts.add_neq_o in H2. +inversion H2. subst. clear H2. +apply IHl. +clear IHl. +unfold adj_set. rewrite <-H5. +rewrite <-H3. assumption. +intro. +assert (InA Vertex.eq (fst_ext e) (a :: l)). +apply HH. assumption. +inversion H4; subst. +elim (n H7). +assumption. +intros. apply HH1. auto. +intros. apply Hneq. right. auto. +auto. +Qed. + +Add Morphism In_graph_edge : In_graph_edge_m. + +Proof. +unfold In_graph_edge;intros x y H g. +fold (eq x y) in H. +split;intro H0;[rewrite <-H|rewrite H];assumption. +Qed. + +(* There cannot exist both an interference and + a preference between two vertices *) +Lemma interf_pref_conflict : forall x y g, +Prefere x y g /\ Interfere x y g -> False. + +Proof. +intros. +unfold Prefere in H. unfold Interfere in H. +destruct H. +apply (simple_graph g x y). +split. +destruct H0. +generalize (proj1 (In_graph_aff_edge_in_AE _ _) H0). intro H1. +destruct H1 as [H1 _]. +unfold aff_edge in H1. +destruct H1. inversion H1. +unfold IE in H0. +rewrite edge_comm in H0. apply (proj1 (edgemap_to_edgeset_charac _ _ _ _ (sym_imap g)) H0). +destruct H. destruct H. generalize (AE_weights _ _ H). simpl. intro. +rewrite H1 in H. +rewrite edge_comm in H. unfold AE in H. apply (proj1 (edgemap_to_edgeset_charac _ _ _ _(sym_pmap g)) H). +generalize (proj1 (In_graph_interf_edge_in_IE _ _) H). intro H1. +destruct H1 as [H1 _]. +unfold interf_edge in H1. inversion H1. +Qed. + +Lemma is_simple_graph : forall e e' g, +In_graph_edge e g -> +In_graph_edge e' g -> +weak_eq e e' -> +eq e e'. + +Proof. +intros. +unfold weak_eq in H1. +destruct H; destruct H0. +generalize (AE_weights _ _ H); intros. +generalize (AE_weights _ _ H0); intros. +destruct H1; destruct H1. +rewrite (edge_eq e). rewrite (edge_eq e'). rewrite H2. rewrite H3. Eq_eq. +rewrite (edge_eq e). rewrite (edge_eq e'). rewrite H2. rewrite H3. Eq_comm_eq. +elim (interf_pref_conflict (fst_ext e') (snd_ext e') g). +unfold Prefere, Interfere; split. +destruct H1; destruct H1. exists N0. +rewrite <-(AE_weights _ _ H). +assert (eq (fst_ext e', snd_ext e', get_weight e) (fst_ext e, snd_ext e, get_weight e)) by Eq_eq. +rewrite H3. rewrite <-(edge_eq e). rewrite In_graph_aff_edge_in_AE in H. destruct H. intuition. +exists N0. rewrite <-(AE_weights _ _ H). +assert (eq (fst_ext e', snd_ext e', get_weight e) (fst_ext e, snd_ext e, get_weight e)) by Eq_comm_eq. +rewrite H3. rewrite <-(edge_eq e). rewrite In_graph_aff_edge_in_AE in H. destruct H. intuition. +rewrite <-(IE_weights _ _ H0). rewrite In_graph_interf_edge_in_IE in H0. destruct H0. rewrite (edge_eq e') in H2. intuition. +elim (interf_pref_conflict (fst_ext e') (snd_ext e') g). +unfold Prefere, Interfere; split. +exists N0. rewrite <-(AE_weights _ _ H0). +rewrite (edge_eq e') in H0. rewrite In_graph_aff_edge_in_AE in H0. destruct H0. auto. +rewrite <-(IE_weights _ _ H). +destruct H1; destruct H1. +assert (eq (fst_ext e', snd_ext e', get_weight e) (fst_ext e, snd_ext e, get_weight e)) by Eq_eq. +rewrite H3. rewrite <-(edge_eq e). rewrite In_graph_interf_edge_in_IE in H. destruct H. auto. +assert (eq (fst_ext e', snd_ext e', get_weight e) (fst_ext e, snd_ext e, get_weight e)) by Eq_comm_eq. +rewrite H3. rewrite <-(edge_eq e). rewrite In_graph_interf_edge_in_IE in H. destruct H. auto. +generalize (IE_weights _ _ H); intros. +generalize (IE_weights _ _ H0); intros. +destruct H1; destruct H1. +rewrite (edge_eq e). rewrite (edge_eq e'). rewrite H2. rewrite H3. Eq_eq. +rewrite (edge_eq e). rewrite (edge_eq e'). rewrite H2. rewrite H3. Eq_comm_eq. +Qed. + +Lemma is_precolored_ext : forall x y g, +Vertex.eq x y -> +is_precolored x g = is_precolored y g. + +Proof. +intros. apply compat_bool_is_precolored. assumption. +Qed. + +(* This module respects the interface *) + +(* A vertex x is in (remove_vertex r g) iff it is in g + and it is different from r*) +Lemma In_remove_vertex : forall x r g, +In_graph x (remove_vertex r g) <-> (In_graph x g /\ ~Vertex.eq x r). + +Proof. +split; intros. +split. apply in_remove_in with (r:=r). auto. +intro. rewrite H0 in H. elim (not_in_remove r g H). +destruct H. apply In_remove; auto. +Qed. + +(* The precolored vertices of (remove_vertex r g) are + the precolored vertices of g, minus r *) +Lemma precolored_remove_vertex : forall r g, +VertexSet.Equal (precolored (remove_vertex r g)) + (VertexSet.remove r (precolored g)). + +Proof. +split; intros. +rewrite precolored_remove_vertex2; auto. +rewrite <-precolored_remove_vertex2; auto. +Qed. + +(* An edge e is in (remove_vertex r g) iff it is in g + and is not incident to r *) +Lemma In_remove_edge : forall e r g, +In_graph_edge e (remove_vertex r g) <-> (In_graph_edge e g /\ ~incident e r). + +Proof. +split; intros. +split. apply in_remove_in_edge with (r:=r). auto. +intro. elim (not_in_remove_edge _ _ _ H0 H). +destruct H. apply In_remove_edge_; auto. +Qed. + +(* Specification of merge *) + +(* A vertex is in (merge e g p q) iff x is in g and + x is not equal to the second endpoint of g *) +Lemma In_merge_vertex : forall g e x p q, +In_graph x (merge e g p q) <-> (In_graph x g /\ ~Vertex.eq x (snd_ext e)). + +Proof. +split; intros. +split. apply (In_merge_in g e x p q H). +intro. rewrite H0 in H. elim (merge_1 e g p q H). +destruct H. apply merge_2; auto. +Qed. + +(* If an interference edge e' is in the graph g then + its redirection from the second endpoint of e + to the first endpoint of e is in (merge e g p q) *) +Lemma In_merge_interf_edge : forall e e' g p q, +In_graph_edge e' g -> +interf_edge e' -> +In_graph_edge (redirect (snd_ext e) (fst_ext e) e') (merge e g p q). + +Proof. +intros. apply merge_3; auto. +Qed. + +(* If a preference edge e' different from e is in the graph g, + and iff the endpoints of its redirection from the second + endpoint of e to the first endpoint of e do not interfere, + then these endpoints are linked with an affinity edge, whose + weight may be different from the one of e' *) +Lemma In_merge_pref_edge : forall e e' g p q, +In_graph_edge e' g -> +~Edge.eq e e' -> +aff_edge e' -> +~Interfere (fst_ext (redirect (snd_ext e) (fst_ext e) e')) + (snd_ext (redirect (snd_ext e) (fst_ext e) e')) + (merge e g p q) -> + Prefere (fst_ext (redirect (snd_ext e) (fst_ext e) e')) + (snd_ext (redirect (snd_ext e) (fst_ext e) e')) + (merge e g p q). + +Proof. +intros. apply merge_5; auto. +Qed. + +(* Inversely, if e' is an edge of (merge e g p q) then there exists + an edge a of g, such that e' is weakly equal to the redirection + of a from the second endpoint of e to the first endpoint of e *) +Lemma In_merge_edge_inv : forall e e' g p q, +In_graph_edge e' (merge e g p q) -> +exists a, In_graph_edge a g /\ weak_eq e' (redirect (snd_ext e) (fst_ext e) a) /\ + same_type a e'. + +Proof. +intros. apply (merge_4 e e' g p q); auto. +Qed. + +(* The precolored vertices of (merge e g p q) are the ones of g, + minus the second endpoint of e *) +Lemma precolored_merge : forall e g p q, +VertexSet.Equal (precolored (merge e g p q)) + (VertexSet.remove (snd_ext e) (precolored g)). + +Proof. +split; intros. +rewrite <-precolored_merge2 in H. auto. +rewrite <-precolored_merge2. auto. +Qed. + +(* Specification of delete_preference_edges *) + +(* A vertex is in (delete_preference_edges r g p) iff it is in g *) +Lemma In_delete_preference_edges_vertex : forall x r g p, +In_graph x (delete_preference_edges r g p) <-> In_graph x g. + +Proof. +unfold In_graph. split; intros. +rewrite <-V_delete_preference_eq in H. auto. +rewrite <-V_delete_preference_eq. auto. +Qed. + +(* The precolored vertices of (delete_preference_edges r g p) + iff it is precolored in g *) +Lemma precolored_delete_preference_edges : forall r g p, +VertexSet.Equal (precolored (delete_preference_edges r g p)) + (precolored g). + +Proof. +intros. apply delete_preference_edges_prec. +Qed. + +(* An edge e is in (delete_preference_edges r g p) iff + if is in g and it is not an affinity edge incident to r *) +Lemma In_delete_preference_edges_edge : forall e r g p, +In_graph_edge e (delete_preference_edges r g p) <-> +(In_graph_edge e g /\ ~(aff_edge e /\ incident e r)). + +Proof. +split; intros. +split. apply (delete_preference_edges_1 _ _ _ _ H). +intro. destruct H0. +assert (~incident e r). +apply (in_delete_preference_not_incident e r g p). +rewrite In_graph_aff_edge_in_AE. split; auto. +elim (H2 H1). +destruct H. destruct H. +rewrite In_graph_aff_edge_in_AE in H. destruct H. +apply delete_preference_edges_2; auto. +right. rewrite <-IE_delete_preference_eq. auto. +Qed. + +Definition interference_adj v g := +adj_set v (imap g). + +Definition preference_adj v g := +adj_set v (pmap g). + +(* Definition of the interference and preference degrees *) +Definition interf_degree g v := VertexSet.cardinal (interference_adj v g). +Definition pref_degree g v := VertexSet.cardinal (preference_adj v g). + +(* Definition of the low-degree function, + returns true iff the interference degree of v in g is strictly lower than K *) +Definition has_low_degree g K v := +if le_lt_dec K (interf_degree g v) then false else true. + +(* Definition of the move-related function, + returns true iff the vertex is move-related *) +Definition move_related g x := negb (VertexSet.is_empty (preference_adj x g)). + +Lemma in_pref_pref : forall x y g, +VertexSet.In x (preference_adj y g) -> +exists w, In_graph_edge (x,y,Some w) g. + +Proof. +intros. +unfold preference_adj in H. +exists N0. +left. unfold AE. +rewrite edge_comm. apply (proj2 (edgemap_to_edgeset_charac _ _ _ _ (sym_pmap g))). assumption. +Qed. + +Lemma pref_in_pref : forall x y g w, +In_graph_edge (x,y,Some w) g -> +VertexSet.In x (preference_adj y g). + +Proof. +intros. +unfold preference_adj. +destruct H. +generalize (AE_weights _ _ H). intro. +unfold AE in H. simpl in H0. rewrite H0 in H. +rewrite edge_comm in H. apply (proj1 (edgemap_to_edgeset_charac _ _ _ _(sym_pmap g)) H). +generalize (proj1 (In_graph_interf_edge_in_IE _ _) H). intro. +destruct H0. +unfold interf_edge in H0. inversion H0. +Qed. + +Lemma in_interf_interf : forall x y g, +VertexSet.In x (interference_adj y g) -> +In_graph_edge (x,y,None) g. + +Proof. +intros. +unfold preference_adj in H. +right. unfold IE. +rewrite edge_comm. apply (proj2 (edgemap_to_edgeset_charac _ _ _ _ (sym_imap g))). assumption. +Qed. + +Lemma interf_in_interf : forall x y g, +In_graph_edge (x,y,None) g -> +VertexSet.In x (interference_adj y g). + +Proof. +intros. +unfold interference_adj. +destruct H. +generalize (AE_weights _ _ H). intro. inversion H0. +rewrite edge_comm in H. apply (proj1 (edgemap_to_edgeset_charac _ _ _ _(sym_imap g)) H). +Qed. + +Lemma compat_interference_adj : forall x y g, +Vertex.eq x y -> +interference_adj x g = interference_adj y g. + +Proof. +intros. +unfold interference_adj. +unfold adj_set. +rewrite (InterfFacts.find_o _ H). reflexivity. +Qed. + +Lemma compat_preference_adj : forall x y g, +Vertex.eq x y -> +preference_adj x g = preference_adj y g. + +Proof. +intros. +unfold preference_adj. +unfold adj_set. +rewrite (InterfFacts.find_o _ H). reflexivity. +Qed. + +Lemma compat_bool_move : forall g, +compat_bool Vertex.eq (move_related g). + +Proof. +intros. +unfold compat_bool. +intros. +unfold move_related. +rewrite (compat_preference_adj _ _ _ H). +reflexivity. +Qed. + +(* characterisation of move relation *) + +Lemma move_related_charac : forall x g, +move_related g x = true -> +exists e, aff_edge e /\ In_graph_edge e g /\ incident e x. + +Proof. +intros. +unfold move_related in H. +case_eq (VertexSet.is_empty (preference_adj x g)); intros. +rewrite H0 in H. inversion H. generalize H0. clear H H0. intro H. +case_eq (VertexSet.choose (preference_adj x g)); intros. +exists (x,e,Some N0). +split. +exists N0. auto. +generalize (VertexSet.choose_1 H0). clear H0. intro H0. +split. +left. unfold AE. apply (proj2 (edgemap_to_edgeset_charac _ _ _ _(sym_pmap g))). +assumption. +left. auto. +generalize (VertexSet.choose_2 H0). clear H0. intro H0. +rewrite (VertexSet.is_empty_1 H0) in H. inversion H. +Qed. + +(* the inversion characterisation of move related *) + +Lemma move_related_charac2 : forall x g e, +aff_edge e -> +In_graph_edge e g -> +incident e x -> +move_related g x = true. + +Proof. +intros. +unfold move_related. +case_eq (VertexSet.is_empty (preference_adj x g)); intros. +generalize (VertexSet.is_empty_2 H2). clear H2. intro H2. +destruct H1. +elim (H2 (snd_ext e)). +apply pref_in_pref with (w:=N0). +rewrite edge_comm. +assert (eq (x,snd_ext e, Some N0) e). +apply eq_ordered_eq. +unfold E.eq; simpl; intuition. apply Regs.eq_refl. +destruct H0. generalize (AE_weights _ _ H0). intro. +rewrite <-H3. apply OptionN_as_OT.eq_refl. +generalize (proj1 (In_graph_interf_edge_in_IE _ _) H0). intros. +destruct H3. +unfold aff_edge in H. rewrite H3 in H. destruct H. inversion H. +rewrite H3. assumption. + +elim (H2 (fst_ext e)). +apply pref_in_pref with (w:=N0). +assert (eq (fst_ext e, x, Some N0) e). +apply eq_ordered_eq. +unfold E.eq; simpl; intuition. apply Regs.eq_refl. +destruct H0. generalize (AE_weights _ _ H0). intro. +rewrite <-H3. apply OptionN_as_OT.eq_refl. +generalize (proj1 (In_graph_interf_edge_in_IE _ _) H0). intros. +destruct H3. +unfold aff_edge in H. rewrite H3 in H. destruct H. inversion H. +rewrite H3. assumption. +auto. +Qed. +Definition WS := (VertexSet.t*VertexSet.t*VertexSet.t*EdgeSet.t)%type. + +Definition get_spillWL (w : WS) := fst (fst (fst w)). +Definition get_freezeWL (w : WS) := snd (fst (fst w)). +Definition get_simplifyWL (w : WS) := snd (fst w). +Definition get_movesWL (w : WS) := snd w. + +Lemma compat_bool_low : forall g palette, +compat_bool Vertex.eq (has_low_degree g palette). + +Proof. +unfold compat_bool. intros. +unfold has_low_degree, interf_degree. +rewrite (compat_interference_adj _ _ _ H). +reflexivity. +Qed. + +Definition WS_properties g K (WL : WS) : Prop := +(forall x, VertexSet.In x (get_spillWL WL) <-> has_low_degree g K x = false /\ In_graph x g /\ ~VertexSet.In x (precolored g)) /\ +(forall x, VertexSet.In x (get_freezeWL WL) <-> has_low_degree g K x = true /\ (move_related g) x = true /\ ~VertexSet.In x (precolored g)) /\ +(forall x, VertexSet.In x (get_simplifyWL WL) <-> has_low_degree g K x = true /\ (move_related g) x = false /\ In_graph x g /\ ~VertexSet.In x (precolored g)) /\ +(forall e, EdgeSet.In e (get_movesWL WL) <-> aff_edge e /\ In_graph_edge e g). + +Definition get_WL g palette := +let not_pre := VertexSet.diff (V g) (precolored g) in +let (low, spill) := VertexSet.partition (has_low_degree g palette) not_pre in +let (free, simp) := VertexSet.partition (move_related g) low in +(spill, free, simp, AE g). + +Module Import RegOTFacts := MyOTFacts Vertex. + +Lemma move_related_in : forall g x, +move_related g x = true -> +In_graph x g. + +Proof. +intros. +generalize (move_related_charac _ _ H). clear H. intro H. +destruct H as [e H]. destruct H. destruct H0. +apply (proj1 (extremities_pmap g x)). +destruct H0. +generalize (AE_weights _ _ H0). intro Hw. +unfold AE in *. +destruct H1. +rewrite (edge_eq e) in H0. +simpl in Hw. rewrite Hw in H0. +generalize (proj1 (edgemap_to_edgeset_charac _ _ _ _(sym_pmap g)) H0). intro H2. +apply (proj2 (InterfFacts.in_find_iff _ _)). +unfold adj_set in H2. +case_eq (VertexMap.find (fst_ext e) (pmap g)); intros; rewrite H3 in H2. +intro Helim. +rewrite (InterfFacts.find_o _ H1) in Helim. +rewrite H3 in Helim. inversion Helim. +elim (VertexSet.empty_1 H2). + +rewrite (edge_eq e) in H0. rewrite edge_comm in H0. +simpl in Hw. rewrite Hw in H0. +generalize (proj1 (edgemap_to_edgeset_charac _ _ _ _(sym_pmap g)) H0). intro H2. +apply (proj2 (InterfFacts.in_find_iff _ _)). +unfold adj_set in H2. +case_eq (VertexMap.find (snd_ext e) (pmap g)); intros; rewrite H3 in H2. +intro Helim. +rewrite (InterfFacts.find_o _ H1) in Helim. +rewrite H3 in Helim. inversion Helim. +elim (VertexSet.empty_1 H2). + +generalize (proj1 (In_graph_interf_edge_in_IE _ _) H0). intros. +destruct H2. unfold aff_edge in H. rewrite H2 in H. inversion H. inversion H4. +Qed. + +Lemma WS_prop_get : forall g palette, +WS_properties g palette (get_WL g palette). + +Proof. +intros. +unfold get_WL. +set (not_pre := VertexSet.diff (V g) (precolored g)) in *. +case_eq (VertexSet.partition (has_low_degree g palette) not_pre). +intros low spill H. +case_eq (VertexSet.partition (move_related g) low). +intros free simp H0. +unfold WS_properties. +unfold get_spillWL. unfold get_simplifyWL. unfold get_freezeWL. unfold get_movesWL. simpl. +assert (VertexSet.Equal low (VertexSet.filter (has_low_degree g palette) not_pre)). +assert (low = fst (VertexSet.partition (has_low_degree g palette) not_pre)). +rewrite H. auto. +rewrite H1. apply VertexSet.partition_1. apply compat_bool_low. +assert (VertexSet.Equal spill (VertexSet.filter (fun x => negb (has_low_degree g palette x)) not_pre)). +assert (spill = snd (VertexSet.partition (has_low_degree g palette) not_pre)). +rewrite H. auto. +rewrite H2. apply VertexSet.partition_2. apply compat_bool_low. +assert (VertexSet.Equal free (VertexSet.filter (move_related g) low)). +assert (free = fst (VertexSet.partition (move_related g) low)). +rewrite H0. auto. +rewrite H3. apply VertexSet.partition_1. apply compat_bool_move. +assert (VertexSet.Equal simp (VertexSet.filter (fun x => negb (move_related g x)) low)). +assert (simp = snd (VertexSet.partition (move_related g) low)). +rewrite H0. auto. +rewrite H4. apply VertexSet.partition_2. apply compat_bool_move. +split; intros. +split; intros. +rewrite H2 in H5. +split. +generalize (VertexSet.filter_2 (compat_not_compat (compat_bool_low _ _)) H5). +destruct (has_low_degree g palette x); intuition. +split. +generalize (VertexSet.filter_1 (compat_not_compat (compat_bool_low _ _)) H5). intro. +apply (VertexSet.diff_1 H6). +generalize (VertexSet.filter_1 (compat_not_compat (compat_bool_low _ _)) H5). intro. +apply (VertexSet.diff_2 H6). +rewrite H2. +apply VertexSet.filter_3. +apply compat_not_compat. apply compat_bool_low. +apply VertexSet.diff_3; intuition. +destruct H5. rewrite H5. auto. + +split; intros. +split; intros. +rewrite H3 in H5. +generalize (VertexSet.filter_1 (compat_bool_move _) H5). intro H6. +generalize (VertexSet.filter_2 (compat_bool_move _) H5). clear H5. intro H5. +rewrite H1 in H6. +generalize (VertexSet.filter_1 (compat_bool_low _ _) H6). intro H7. +generalize (VertexSet.filter_2 (compat_bool_low _ _) H6). clear H6. intro H6. +generalize (VertexSet.diff_2 H7). intuition. + +rewrite H3. apply VertexSet.filter_3. +apply compat_bool_move. +rewrite H1. +apply VertexSet.filter_3. +apply compat_bool_low. +apply VertexSet.diff_3. +apply move_related_in. intuition. +intuition. +intuition. +intuition. + +split;intros. +split;intros. +rewrite H4 in H5. +generalize (VertexSet.filter_1 (compat_not_compat (compat_bool_move _)) H5). intro H6. +generalize (VertexSet.filter_2 (compat_not_compat (compat_bool_move _)) H5). clear H5. intro H5. +rewrite H1 in H6. +generalize (VertexSet.filter_1 (compat_bool_low _ _) H6). intro H7. +generalize (VertexSet.filter_2 (compat_bool_low _ _) H6). clear H6. intro H6. +generalize (VertexSet.diff_2 H7). intuition. +destruct (move_related g x); intuition. +apply (VertexSet.diff_1 H7). + +rewrite H4. +apply VertexSet.filter_3. +apply compat_not_compat. apply compat_bool_move. +rewrite H1. apply VertexSet.filter_3. +apply compat_bool_low. +apply VertexSet.diff_3. +intuition. +intuition. +intuition. +destruct (move_related g x); intuition. +exact (In_graph_aff_edge_in_AE _ _). +Qed. + +Definition not_incident_edges x s g := +VertexSet.fold (fun y s' => EdgeSet.remove (x,y,Some N0) s') + (adj_set x (pmap g)) + s. + +Lemma not_incident_edges_1 : forall x e s g, +(forall y, EdgeSet.In y s -> aff_edge y /\ In_graph_edge y g) -> +(EdgeSet.In e (not_incident_edges x s g) <-> + EdgeSet.In e s /\ ~incident e x). + +Proof. +split; intros. +unfold not_incident_edges in H0. +rewrite VertexSet.fold_1 in H0. +assert (EdgeSet.In e s) as Hin. +induction (VertexSet.elements (adj_set x (pmap g))). simpl in H0. +assumption. +rewrite MEdgeFacts.fold_left_assoc in H0. +apply IHl. +apply (EdgeSet.remove_3 H0). + +intros. +split; intros. +apply EdgeSet.remove_2. intro. +generalize (EdgeSet.remove_3 H1). intro. +elim (EdgeSet.remove_1 H2 H3). +apply EdgeSet.remove_2. intro. +elim (EdgeSet.remove_1 H2 H1). +apply (EdgeSet.remove_3 (EdgeSet.remove_3 H1)). +apply EdgeSet.remove_2. intro. +generalize (EdgeSet.remove_3 H1). intro. +elim (EdgeSet.remove_1 H2 H3). +apply EdgeSet.remove_2. intro. +elim (EdgeSet.remove_1 H2 H1). +apply (EdgeSet.remove_3 (EdgeSet.remove_3 H1)). +intros. apply RegRegProps.Equal_remove. assumption. + +split. +assumption. + +generalize VertexSet.elements_1. intro HH. +intro Helim. destruct Helim. +generalize (HH (adj_set x (pmap g)) (snd_ext e)). clear HH. intro HH. +induction (VertexSet.elements (adj_set x (pmap g))). simpl in H0. + +assert (VertexSet.In (snd_ext e) (adj_set x (pmap g))). +apply (proj1 (edgemap_to_edgeset_charac _ _ _ (Some N0) (sym_pmap g))). +assert (eq e (x, snd_ext e, Some N0)). +assert (get_weight e = Some N0). + +generalize (H _ H0). intro. +destruct H2. unfold In_graph_edge in H3. +destruct H3. +apply (AE_weights g). assumption. +generalize (IE_weights g _ H3). intro. +destruct H2. congruence. +rewrite <-H2. Eq_eq. apply Regs.eq_refl. +rewrite <-H2. + +generalize (H _ H0). intro. +destruct H3. destruct H4. +assumption. +generalize (IE_weights g _ H4). intro. +destruct H3. congruence. + +generalize (HH H2). intro. inversion H3. + +rewrite MEdgeFacts.fold_left_assoc in H0. +apply IHl. +apply (EdgeSet.remove_3 H0). + +intro. generalize (HH H2). clear HH H2. intro H2. +inversion H2; subst. +assert (eq e (x, a, Some N0)). +rewrite (edge_eq e). +apply eq_ordered_eq. +constructor. simpl. split; intuition. +simpl. + +generalize (H _ Hin). intro H5. +destruct H5. +destruct H5. +rewrite (AE_weights g _ H5). apply OptionN_as_OT.eq_refl. +generalize (IE_weights g _ H5). intro. destruct H3. congruence. +elim (EdgeSet.remove_1 (eq_sym H3) H0). +assumption. + +intros. +split; intros. +apply EdgeSet.remove_2. intro. +generalize (EdgeSet.remove_3 H2). intro. +elim (EdgeSet.remove_1 H3 H4). +apply EdgeSet.remove_2. intro. +elim (EdgeSet.remove_1 H3 H2). +apply (EdgeSet.remove_3 (EdgeSet.remove_3 H2)). +apply EdgeSet.remove_2. intro. +generalize (EdgeSet.remove_3 H2). intro. +elim (EdgeSet.remove_1 H3 H4). +apply EdgeSet.remove_2. intro. +elim (EdgeSet.remove_1 H3 H2). +apply (EdgeSet.remove_3 (EdgeSet.remove_3 H2)). +intros. apply RegRegProps.Equal_remove. assumption. + +generalize (HH (adj_set x (pmap g)) (fst_ext e)). clear HH. intro HH. +induction (VertexSet.elements (adj_set x (pmap g))). simpl in H0. + +assert (VertexSet.In (fst_ext e) (adj_set x (pmap g))). +apply (proj1 (edgemap_to_edgeset_charac _ _ _ (Some N0) (sym_pmap g))). +assert (eq e (x, fst_ext e, Some N0)). +assert (get_weight e = Some N0). + +generalize (H _ H0). intro. +destruct H2. unfold In_graph_edge in H3. +destruct H3. +apply (AE_weights g). assumption. +generalize (IE_weights g _ H3). intro. +destruct H2. congruence. +rewrite <-H2. Eq_comm_eq. apply Regs.eq_refl. +rewrite <-H2. + +generalize (H _ H0). intro. +destruct H3. destruct H4. +assumption. +generalize (IE_weights g _ H4). intro. +destruct H3. congruence. + +generalize (HH H2). intro. inversion H3. + +rewrite MEdgeFacts.fold_left_assoc in H0. +apply IHl. +apply (EdgeSet.remove_3 H0). + +intro. generalize (HH H2). clear HH H2. intro H2. +inversion H2; subst. +assert (eq e (x, a, Some N0)). +rewrite (edge_eq e). +rewrite edge_comm. apply eq_ordered_eq. +constructor. simpl. split; intuition. +simpl. + +generalize (H _ Hin). intro H5. +destruct H5. +destruct H5. +rewrite (AE_weights g _ H5). apply OptionN_as_OT.eq_refl. +generalize (IE_weights g _ H5). intro. destruct H3. congruence. +elim (EdgeSet.remove_1 (eq_sym H3) H0). +assumption. + +intros. +split; intros. +apply EdgeSet.remove_2. intro. +generalize (EdgeSet.remove_3 H2). intro. +elim (EdgeSet.remove_1 H3 H4). +apply EdgeSet.remove_2. intro. +elim (EdgeSet.remove_1 H3 H2). +apply (EdgeSet.remove_3 (EdgeSet.remove_3 H2)). +apply EdgeSet.remove_2. intro. +generalize (EdgeSet.remove_3 H2). intro. +elim (EdgeSet.remove_1 H3 H4). +apply EdgeSet.remove_2. intro. +elim (EdgeSet.remove_1 H3 H2). +apply (EdgeSet.remove_3 (EdgeSet.remove_3 H2)). +intros. apply RegRegProps.Equal_remove. assumption. + +unfold not_incident_edges. +rewrite VertexSet.fold_1. +induction (VertexSet.elements (adj_set x (pmap g))). simpl. +intuition. +rewrite MEdgeFacts.fold_left_assoc. +apply EdgeSet.remove_2. +destruct H0. +intro H2. elim H1. +destruct (eq_charac _ _ H2); destruct H3; change_rewrite. +left. auto. +right. auto. +assumption. + +intros. +split; intros. +apply EdgeSet.remove_2. intro. +generalize (EdgeSet.remove_3 H1). intro. +elim (EdgeSet.remove_1 H2 H3). +apply EdgeSet.remove_2. intro. +elim (EdgeSet.remove_1 H2 H1). +apply (EdgeSet.remove_3 (EdgeSet.remove_3 H1)). +apply EdgeSet.remove_2. intro. +generalize (EdgeSet.remove_3 H1). intro. +elim (EdgeSet.remove_1 H2 H3). +apply EdgeSet.remove_2. intro. +elim (EdgeSet.remove_1 H2 H1). +apply (EdgeSet.remove_3 (EdgeSet.remove_3 H1)). +intros. apply RegRegProps.Equal_remove. assumption. +Qed. + +Definition not_incident_merge x s es := +VertexSet.fold (fun y s' => EdgeSet.remove (x,y,Some N0) s') + es s. + +Lemma not_incident_merge_1 : forall x s es e, +EdgeSet.In e (not_incident_merge x es s) -> +EdgeSet.In e es. + +Proof. +unfold not_incident_merge; intros. +set (f := (fun (y : VertexSet.elt) (s' : EdgeSet.t) => + EdgeSet.remove (x, y, Some 0%N) s')) in *. +rewrite VertexSet.fold_1 in H. +set (f' := fun a e => f e a) in *. +induction (VertexSet.elements s). simpl in H. assumption. +rewrite MEdgeFacts.fold_left_assoc in H. +set (tmp := fold_left f' l es) in H. +unfold f' in H. unfold f in H. +apply (IHl (EdgeSet.remove_3 H)). + +unfold f'. unfold f. intros. +split; intros. +apply EdgeSet.remove_2. +intro H1. elim (EdgeSet.remove_1 H1 (EdgeSet.remove_3 H0)). +apply EdgeSet.remove_2. +intro H1. elim (EdgeSet.remove_1 H1 H0). +apply (EdgeSet.remove_3 (EdgeSet.remove_3 H0)). +apply EdgeSet.remove_2. +intro H1. elim (EdgeSet.remove_1 H1 (EdgeSet.remove_3 H0)). +apply EdgeSet.remove_2. +intro H1. elim (EdgeSet.remove_1 H1 H0). +apply (EdgeSet.remove_3 (EdgeSet.remove_3 H0)). + +intros. +unfold f'. unfold f. +apply RegRegProps.Equal_remove. auto. +Qed. + + +Definition AE_merge_up newadj adj adjsnd e es := +let diff1 := VertexSet.diff adj newadj in +let diff2 := VertexSet.diff newadj adj in +let new_es := not_incident_merge (snd_ext e) + (not_incident_merge (fst_ext e) es diff1) + adjsnd in +VertexSet.fold (fun y s'' => EdgeSet.add (fst_ext e,y,Some N0) s'') + diff2 new_es. + +Lemma not_incident_merge_2 : forall x y s s', +EdgeSet.In (x,y,Some N0) (not_incident_merge x s s') -> +~VertexSet.In y s'. + +Proof. +intros. +unfold not_incident_merge in H. +set (f := (fun (y : VertexSet.elt) (s' : EdgeSet.t) => + EdgeSet.remove (x, y, Some 0%N) s')) in *. +rewrite VertexSet.fold_1 in H. +set (f' := fun a e => f e a) in *. +generalize VertexSet.elements_1. intro HH. +generalize (HH s' y). clear HH. intro HH. +induction (VertexSet.elements s'). simpl in H. +intro. generalize (HH H0). intro H1. inversion H1. +rewrite MEdgeFacts.fold_left_assoc in H. +set (tmp := fold_left f' l s) in *. +unfold f' in H. unfold f in H. +destruct (Vertex.eq_dec y a). +cut (eq (x,y,Some N0) (x,a, Some N0)). intro H0. +elim (EdgeSet.remove_1 (Edge.eq_sym H0) H). +apply eq_ordered_eq. constructor;simpl; try split; intuition. +apply OptionN_as_OT.eq_refl. +apply IHl. apply (EdgeSet.remove_3 H). +intro. generalize (HH H0). intro H1. +inversion H1; subst. +elim (n H3). +auto. + +unfold f'. unfold f. intros. +split; intros. +apply EdgeSet.remove_2. +intro H5. elim (EdgeSet.remove_1 H5 (EdgeSet.remove_3 H0)). +apply EdgeSet.remove_2. +intro H5. elim (EdgeSet.remove_1 H5 H0). +apply (EdgeSet.remove_3 (EdgeSet.remove_3 H0)). + +apply EdgeSet.remove_2. +intro H5. elim (EdgeSet.remove_1 H5 (EdgeSet.remove_3 H0)). +apply EdgeSet.remove_2. +intro H5. elim (EdgeSet.remove_1 H5 H0). +apply (EdgeSet.remove_3 (EdgeSet.remove_3 H0)). + +unfold f'. unfold f. intros. +apply RegRegProps.Equal_remove. auto. +Qed. + +Lemma AE_merge_wl_1 : forall x y e g, +aff_edge e -> +In_graph_edge e g -> +(VertexSet.In x (adj_set y (pmap_merge e g (imap_merge e g))) <-> +(VertexSet.In x (adj_set y (map_merge e (pmap g))) /\ +~VertexSet.In x (adj_set y (imap_merge e g)))). + +Proof. +intros x y e g p q. split; intros. +split. +apply pmap_merge_sub. assumption. +intro. apply (simple_graph (merge e g q p) x y). intuition. +destruct H. +unfold merge. simpl. apply resolve_conflicts_map_5; auto. +Qed. + +Lemma not_in_diff_equiv : forall x s s', +~VertexSet.In x (VertexSet.diff s s') -> +VertexSet.In x s' \/ ~VertexSet.In x s. + +Proof. +intros. destruct (Props.In_dec x s). + destruct (Props.In_dec x s'). +left. auto. +elim H. apply VertexSet.diff_3; auto. +right. auto. +Qed. + +Lemma AE_aux_2 : forall e g s p q, +(forall a, EdgeSet.In a s <-> aff_edge a /\ In_graph_edge a g) -> +(forall b, EdgeSet.In b (AE_merge_up (preference_adj (fst_ext e) (merge e g p q)) + (preference_adj (fst_ext e) g) + (preference_adj (snd_ext e) g) + e s) + <-> + (~incident b (snd_ext e) /\ + ((EdgeSet.In b s /\ ~Interfere (fst_ext b) (snd_ext b) (merge e g p q)) \/ + EdgeSet.In b (VertexSet.fold + (fun y s' => EdgeSet.add (fst_ext e, y, Some N0) s') + (preference_adj (fst_ext e) (merge e g p q)) + EdgeSet.empty)))). + +Proof. +intros e g s p q; split; intros. +unfold AE_merge_up in H0. +set (f := (fun (y : VertexSet.elt) (s'' : EdgeSet.t) => + EdgeSet.add (fst_ext e, y, Some 0%N) s'')) in *. +set (diff1 := (VertexSet.diff (preference_adj (fst_ext e) (merge e g p q)) + (preference_adj (fst_ext e) g))) in *. +set (diff2 := (VertexSet.diff (preference_adj (fst_ext e) g) + (preference_adj (fst_ext e) (merge e g p q)))) in *. +cut (get_weight b = Some N0). intro Hw. +cut (~incident b (snd_ext e)). intro Hcut. +split. +assumption. +rewrite VertexSet.fold_1 in H0. +set (f' := fun a e => f e a) in *. +generalize VertexSet.elements_2. intro HH. +generalize (HH diff1). clear HH. intro HH. +induction (VertexSet.elements diff1). simpl in H0. +left. split. +apply (not_incident_merge_1 _ _ _ _ (not_incident_merge_1 _ _ _ _ H0)). +intro H1. unfold Interfere in H1. +generalize (merge_4 e (fst_ext b, snd_ext b, None) _ p q H1). intro H2. +destruct H2. destruct H2. +unfold redirect in H3. +destruct (OTFacts.eq_dec (fst_ext x) (snd_ext e)). destruct H3. +assert (eq (fst_ext b, snd_ext b, None) (fst_ext e, snd_ext x, get_weight x)). +unfold weak_eq in H3. change_rewrite. destruct H3. +apply eq_ordered_eq; split; simpl; auto. +unfold same_type in H4. destruct H4. +destruct H4. destruct H5. simpl in H5. inversion H5. +destruct H4. rewrite <-H4. apply OptionN_as_OT.eq_refl. +destruct H3. +unfold same_type in H4. destruct H4. +destruct H4. destruct H6. simpl in H6. inversion H6. +destruct H4. rewrite <-H4. rewrite edge_comm. apply eq_ordered_eq. +split; auto. simpl. apply OptionN_as_OT.eq_refl. +generalize H5. clear H3 H4 H5. intro H3. +destruct (eq_charac _ _ H3); destruct H4; change_rewrite. +generalize (not_incident_merge_1 _ _ _ _ H0). intro H6. +cut (eq b (fst_ext e, snd_ext x, Some N0)). intro H7. +rewrite H7 in H6. clear H7. +generalize (not_incident_merge_2 _ _ _ _ H6). intro H7. +unfold diff2 in H7. +generalize (not_in_diff_equiv _ _ _ H7). clear H7. intro H7. destruct H7. +rewrite <-H5 in H7. generalize (sym_pmap_merge_map e g q p _ _ H7). intro H8. +rewrite <-H4 in H8. generalize (sym_pmap_merge_map e g q p _ _ H8). intro H9. +elim (simple_graph (merge e g p q) (snd_ext b) (fst_ext b)). split. +apply (proj1 (edgemap_to_edgeset_charac _ _ _ None (sym_imap_merge_map e g q p))). +destruct H1. +generalize (AE_weights _ _ H1). simpl. congruence. +assumption. +assumption. +rewrite <-H5 in H7. +elim H7. apply (sym_pmap g). +rewrite <-H4. +cut (In_graph_edge b g). +intro H8. +destruct H8. +apply (proj1 (edgemap_to_edgeset_charac _ _ _ (Some N0) (sym_pmap g))). +cut (eq b (snd_ext b, fst_ext b, Some N0)). intro H9. +rewrite <-H9. assumption. +rewrite edge_comm. rewrite (edge_eq b). change_rewrite. +rewrite Hw. apply eq_refl. +generalize (IE_weights _ _ H8). congruence. +generalize (proj1 (H _) (not_incident_merge_1 _ _ _ _ (not_incident_merge_1 _ _ _ _ H0))). +intuition. +apply eq_ordered_eq. +constructor; simpl. split. auto. auto. +fold (get_weight b). rewrite Hw. apply OptionN_as_OT.eq_refl. + +generalize (not_incident_merge_1 _ _ _ _ H0). intro H6. +cut (eq b (fst_ext e, snd_ext x, Some N0)). intro H7. +rewrite H7 in H6. clear H7. +generalize (not_incident_merge_2 _ _ _ _ H6). intro H7. +unfold diff2 in H7. +generalize (not_in_diff_equiv _ _ _ H7). clear H7. intro H7. +destruct H7. +rewrite <-H4 in H7. generalize (sym_pmap_merge_map e g q p _ _ H7). intro H8. +rewrite <-H5 in H8. generalize (sym_pmap_merge_map e g q p _ _ H8). intro H9. +elim (simple_graph (merge e g p q) (snd_ext b) (fst_ext b)). split. +apply (proj1 (edgemap_to_edgeset_charac _ _ _ None (sym_imap_merge_map e g q p))). +destruct H1. +generalize (AE_weights _ _ H1). simpl. congruence. +assumption. +assumption. +rewrite <-H4 in H7. +elim H7. apply (sym_pmap g). +rewrite <-H5. +cut (In_graph_edge b g). +intro H8. +destruct H8. +apply (proj1 (edgemap_to_edgeset_charac _ _ _ (Some N0) (sym_pmap g))). +cut (eq b (snd_ext b, fst_ext b, Some N0)). intro H9. +rewrite edge_comm. rewrite <-H9. assumption. +rewrite edge_comm. rewrite (edge_eq b). change_rewrite. +rewrite Hw. apply eq_refl. +generalize (IE_weights _ _ H8). congruence. +generalize (proj1 (H _) (not_incident_merge_1 _ _ _ _ (not_incident_merge_1 _ _ _ _ H0))). +intuition. +rewrite edge_comm. apply eq_ordered_eq. +constructor; simpl. split. auto. auto. +fold (get_weight b). rewrite Hw. apply OptionN_as_OT.eq_refl. + +destruct (OTFacts.eq_dec (snd_ext x) (snd_ext e)). +assert (eq (fst_ext b, snd_ext b, None) (fst_ext x, fst_ext e, get_weight x)). +destruct H3. unfold same_type in H4. destruct H4; destruct H4. +destruct H5. simpl in H5. inversion H5. rewrite <-H4. +unfold weak_eq in H3. change_rewrite. +destruct H3; destruct H3. +apply eq_ordered_eq. split; auto. simpl. apply OptionN_as_OT.eq_refl. +rewrite edge_comm. apply eq_ordered_eq. split; auto. simpl. apply OptionN_as_OT.eq_refl. +generalize H4. clear H3 H4. intro H3. +destruct (eq_charac _ _ H3); change_rewrite. destruct H4. +generalize (not_incident_merge_1 _ _ _ _ H0). intro H6. +cut (eq b (fst_ext e, fst_ext x, Some N0)). intro H7. +rewrite H7 in H6. clear H7. +generalize (not_incident_merge_2 _ _ _ _ H6). intro H7. +unfold diff2 in H7. +generalize (not_in_diff_equiv _ _ _ H7). clear H7. intro H7. +destruct H7. +rewrite <-H4 in H7. generalize (sym_pmap_merge_map e g q p _ _ H7). intro H8. +rewrite <-H5 in H8. generalize (sym_pmap_merge_map e g q p _ _ H8). intro H9. +elim (simple_graph (merge e g p q) (snd_ext b) (fst_ext b)). split. +apply (proj1 (edgemap_to_edgeset_charac _ _ _ None (sym_imap_merge_map e g q p))). +destruct H1. +generalize (AE_weights _ _ H1). simpl. congruence. +assumption. +assumption. +rewrite <-H4 in H7. +elim H7. apply (sym_pmap g). +rewrite <-H5. +cut (In_graph_edge b g). +intro H8. +destruct H8. +apply (proj1 (edgemap_to_edgeset_charac _ _ _ (Some N0) (sym_pmap g))). +cut (eq b (fst_ext b, snd_ext b, Some N0)). intro H9. +rewrite <-H9. assumption. +rewrite (edge_eq b). change_rewrite. +rewrite Hw. apply eq_refl. +generalize (IE_weights _ _ H8). congruence. +generalize (proj1 (H _) (not_incident_merge_1 _ _ _ _ (not_incident_merge_1 _ _ _ _ H0))). +intuition. +rewrite edge_comm. apply eq_ordered_eq. +constructor; simpl. split. auto. auto. +fold (get_weight b). rewrite Hw. apply OptionN_as_OT.eq_refl. + +destruct H4. +generalize (not_incident_merge_1 _ _ _ _ H0). intro H6. +cut (eq b (fst_ext e, fst_ext x, Some N0)). intro H7. +rewrite H7 in H6. clear H7. +generalize (not_incident_merge_2 _ _ _ _ H6). intro H7. +unfold diff2 in H7. +generalize (not_in_diff_equiv _ _ _ H7). clear H7. intro H7. +destruct H7. +rewrite <-H5 in H7. generalize (sym_pmap_merge_map e g q p _ _ H7). intro H8. +rewrite <-H4 in H8. generalize (sym_pmap_merge_map e g q p _ _ H8). intro H9. +elim (simple_graph (merge e g p q) (snd_ext b) (fst_ext b)). split. +apply (proj1 (edgemap_to_edgeset_charac _ _ _ None (sym_imap_merge_map e g q p))). +destruct H1. +generalize (AE_weights _ _ H1). simpl. congruence. +assumption. +assumption. +rewrite <-H5 in H7. +elim H7. apply (sym_pmap g). +rewrite <-H4. +cut (In_graph_edge b g). +intro H8. +destruct H8. +apply (proj1 (edgemap_to_edgeset_charac _ _ _ (Some N0) (sym_pmap g))). +cut (eq b (fst_ext b, snd_ext b, Some N0)). intro H9. +rewrite edge_comm. rewrite <-H9. assumption. +rewrite (edge_eq b). change_rewrite. +rewrite Hw. apply eq_refl. +generalize (IE_weights _ _ H8). congruence. +generalize (proj1 (H _) (not_incident_merge_1 _ _ _ _ (not_incident_merge_1 _ _ _ _ H0))). +intuition. +apply eq_ordered_eq. +constructor; simpl. split. auto. auto. +fold (get_weight b). rewrite Hw. apply OptionN_as_OT.eq_refl. + +assert (eq (fst_ext b, snd_ext b, None) x). +destruct H3. unfold same_type in H4. destruct H4; destruct H4. +destruct H5. simpl in H5. inversion H5. +rewrite <-H4. +unfold weak_eq in H3; change_rewrite. destruct H3; destruct H3. +apply eq_ordered_eq; split; auto. apply OptionN_as_OT.eq_refl. +rewrite edge_comm. apply eq_ordered_eq; split; auto. apply OptionN_as_OT.eq_refl. +generalize H4. clear H3 H4. intro H3. +elim (simple_graph g (fst_ext b) (snd_ext b)). +split. +apply (proj1 (edgemap_to_edgeset_charac _ _ _ None (sym_imap g))). +rewrite edge_comm. rewrite H3. +destruct H2. +rewrite <-H3 in H2. generalize (AE_weights _ _ H2). simpl. congruence. +assumption. +apply (proj1 (edgemap_to_edgeset_charac _ _ _ (Some N0) (sym_pmap g))). +assert (In_graph_edge b g). +generalize (proj1 (H _) (not_incident_merge_1 _ _ _ _ (not_incident_merge_1 _ _ _ _ H0))). +intuition. +destruct H4. +assert (eq b (fst_ext b, snd_ext b, Some N0)). +rewrite (edge_eq b). change_rewrite. +rewrite Hw. apply eq_refl. +rewrite edge_comm. rewrite <-H5. assumption. +generalize (IE_weights _ _ H4). congruence. +rewrite MEdgeFacts.fold_left_assoc in H0. +set (set := (not_incident_merge (snd_ext e) + (not_incident_merge (fst_ext e) s diff2) + (preference_adj (snd_ext e) g))) in *. +set (tmp := fold_left f' l set) in *. +unfold f' in H0. unfold f in H0. +destruct (proj1 (RegRegProps.Dec.F.add_iff _ _ _) H0). +fold (eq (fst_ext e, a, Some N0) b) in H1. +right. +rewrite <-H1. +rewrite VertexSet.fold_1. +fold f'. +set (s' := preference_adj (fst_ext e) (merge e g p q)) in *. +generalize VertexSet.elements_1. intro HHH. +generalize (HHH s' a). clear HHH. intro HHH. +induction (VertexSet.elements s'). simpl. +assert (InA Vertex.eq a nil). +apply HHH. +unfold s'. +assert (VertexSet.In a diff1). +apply HH. left. intuition. +apply (VertexSet.diff_1 H2). inversion H2. +rewrite MEdgeFacts.fold_left_assoc. +set (tmp' := fold_left f' l0 EdgeSet.empty) in *. +unfold f'. unfold f. +destruct (Vertex.eq_dec a a0). +apply EdgeSet.add_1. +apply eq_ordered_eq. constructor;simpl. +split; intuition. +apply OptionN_as_OT.eq_refl. +apply EdgeSet.add_2. +apply IHl0. intros. +generalize (HHH H2). intro. +inversion H3; subst. +elim (n H5). +auto. + +unfold f'. unfold f. intros. +apply RegRegProps.add_add. + +unfold f'. unfold f. intros. +apply RegRegProps.Dec.F.add_m. apply eq_refl. auto. + +apply IHl. +assumption. +intros. +apply HH. +right. auto. + +unfold f'. unfold f. intros. +apply RegRegProps.add_add. + +unfold f'. unfold f. intros. +apply RegRegProps.Dec.F.add_m. apply eq_refl. auto. + +rewrite VertexSet.fold_1 in H0. +set (f' := fun a e => f e a) in *. +generalize VertexSet.elements_2. intro HH. +generalize (HH diff1). clear HH. intro HH. +induction (VertexSet.elements diff1). simpl in H0. +intro Hinc. destruct Hinc. +assert (eq b (snd_ext e, snd_ext b, Some N0)). +rewrite (edge_eq b). change_rewrite. rewrite Hw. +apply eq_ordered_eq. +constructor;simpl. split;intuition. +apply OptionN_as_OT.eq_refl. + +rewrite H2 in H0. +generalize (not_incident_merge_2 _ _ _ _ H0). intro H3. +elim H3. +apply (proj1 (edgemap_to_edgeset_charac _ _ _ (Some N0) (sym_pmap g))). rewrite <-H2. +assert (In_graph_edge b g). +generalize (proj1 (H _) (not_incident_merge_1 _ _ _ _ (not_incident_merge_1 _ _ _ _ H0))). +intro. destruct H4. rewrite <-H2 in H5. assumption. +destruct H4. +assumption. +generalize (IE_weights _ _ H4). congruence. + +assert (eq b (fst_ext b, snd_ext e, Some N0)). +rewrite (edge_eq b). change_rewrite. rewrite Hw. +apply eq_ordered_eq. +constructor;simpl. split; intuition. +apply OptionN_as_OT.eq_refl. + +rewrite H2 in H0. rewrite edge_comm in H0. +generalize (not_incident_merge_2 _ _ _ _ H0). intro H3. +elim H3. +apply (proj1 (edgemap_to_edgeset_charac _ _ _ (Some N0) (sym_pmap g))). rewrite edge_comm. rewrite <-H2. +assert (In_graph_edge b g). +generalize (proj1 (H _) (not_incident_merge_1 _ _ _ _ (not_incident_merge_1 _ _ _ _ H0))). +intro. destruct H4. rewrite edge_comm in H2. rewrite <-H2 in H5. assumption. +destruct H4. +assumption. +generalize (IE_weights _ _ H4). congruence. + +rewrite MEdgeFacts.fold_left_assoc in H0. +set (set := (not_incident_merge (snd_ext e) + (not_incident_merge (fst_ext e) s diff2) + (preference_adj (snd_ext e) g))) in *. +set (tmp := fold_left f' l set) in *. +unfold f' in H0. unfold f in H0. +destruct (proj1 (RegRegProps.Dec.F.add_iff _ _ _) H0). +fold (eq (fst_ext e, a, Some N0) b) in H1. +intro. destruct H2. +destruct (eq_charac _ _ H1); change_rewrite. destruct H3. +elim (In_graph_edge_diff_ext _ _ p). +apply Vertex.eq_trans with (y := fst_ext b); auto. +destruct H3. +assert (VertexSet.In a diff1). apply HH. +left. intuition. +unfold diff1 in H5. generalize (VertexSet.diff_2 H5). clear H5. intro H5. +elim H5. +rewrite H4. rewrite <-H2. +apply (proj1 (edgemap_to_edgeset_charac _ _ _ (Some N0) (sym_pmap g))). +generalize p. intro Hin. +destruct Hin. +assert (eq (fst_ext e, snd_ext e, Some N0) e). +apply eq_ordered_eq. +constructor. simpl. split; apply Regs.eq_refl. +simpl. fold (get_weight e). rewrite (AE_weights _ _ H6). +apply OptionN_as_OT.eq_refl. +rewrite H7. assumption. +generalize (IE_weights _ _ H6). generalize q. intro Haff. destruct Haff. congruence. + +destruct (eq_charac _ _ H1); change_rewrite. destruct H3. +assert (VertexSet.In a diff1). apply HH. +left. intuition. +unfold diff1 in H5. generalize (VertexSet.diff_2 H5). clear H5. intro H5. +elim H5. +rewrite H4. rewrite <-H2. +apply (proj1 (edgemap_to_edgeset_charac _ _ _ (Some N0) (sym_pmap g))). +generalize p. intro Hin. +destruct Hin. +assert (eq (fst_ext e, snd_ext e, Some N0) e). +apply eq_ordered_eq. +constructor. simpl. split; apply Regs.eq_refl. +simpl. fold (get_weight e). rewrite (AE_weights _ _ H6). +apply OptionN_as_OT.eq_refl. +rewrite H7. assumption. +generalize (IE_weights _ _ H6). generalize q. intro Haff. destruct Haff. congruence. +elim (In_graph_edge_diff_ext _ _ p). +apply Vertex.eq_trans with (y := snd_ext b); auto. +destruct H3. auto. +apply IHl. auto. +intros. apply HH. right. auto. + +unfold f'. unfold f. intros. +apply RegRegProps.add_add. + +unfold f'. unfold f. intros. +apply RegRegProps.Dec.F.add_m. apply eq_refl. auto. + +rewrite VertexSet.fold_1 in H0. +set (f' := fun a e => f e a) in *. +induction (VertexSet.elements diff1). simpl in H0. +assert (aff_edge b /\ In_graph_edge b g). +rewrite <-H. +apply (not_incident_merge_1 _ _ _ _ (not_incident_merge_1 _ _ _ _ H0)). +destruct H1. +destruct H2. +apply (AE_weights _ _ H2). +destruct H1. generalize (IE_weights _ _ H2). congruence. + +rewrite MEdgeFacts.fold_left_assoc in H0. +set (set := (not_incident_merge (snd_ext e) + (not_incident_merge (fst_ext e) s diff2) + (preference_adj (snd_ext e) g))) in *. +set (tmp := fold_left f' l set) in *. +unfold f' in H0. unfold f in H0. +destruct (proj1 (RegRegProps.Dec.F.add_iff _ _ _) H0). +fold (eq (fst_ext e, a, Some N0) b) in H1. +rewrite <-H1. auto. +apply IHl. auto. + +unfold f'. unfold f. intros. +apply RegRegProps.add_add. + +unfold f'. unfold f. intros. +apply RegRegProps.Dec.F.add_m. apply eq_refl. auto. + +destruct H0. +destruct H1. +destruct H1. +unfold AE_merge_up. +set (f := (fun (y : VertexSet.elt) (s'' : EdgeSet.t) => + EdgeSet.add (fst_ext e, y, Some 0%N) s'')) in *. +set (diff1 := (VertexSet.diff (preference_adj (fst_ext e) (merge e g p q)) + (preference_adj (fst_ext e) g))) in *. +set (diff2 := (VertexSet.diff (preference_adj (fst_ext e) g) + (preference_adj (fst_ext e) (merge e g p q)))) in *. +cut (EdgeSet.In b (not_incident_merge (snd_ext e) + (not_incident_merge (fst_ext e) s diff2) + (preference_adj (snd_ext e) g))). intro H3. +rewrite VertexSet.fold_1. +set (f' := fun a e => f e a) in *. +induction (VertexSet.elements diff1). simpl. assumption. +rewrite MEdgeFacts.fold_left_assoc. +set (set := (fold_left f' l + (not_incident_merge (snd_ext e) + (not_incident_merge (fst_ext e) s diff2) + (preference_adj (snd_ext e) g)))) in *. +unfold f'. unfold f. +apply EdgeSet.add_2. assumption. + +unfold f'. unfold f. intros. +apply RegRegProps.add_add. + +unfold f'. unfold f. intros. +apply RegRegProps.Dec.F.add_m. apply eq_refl. auto. + +cut (EdgeSet.In b (not_incident_merge (fst_ext e) s diff2)). intro. +set (set := (not_incident_merge (fst_ext e) s diff2)) in *. +unfold not_incident_merge. +set (h := (fun (y : VertexSet.elt) (s' : EdgeSet.t) => + EdgeSet.remove (snd_ext e, y, Some 0%N) s')) in *. +set (adjsnd := preference_adj (snd_ext e) g) in *. +rewrite VertexSet.fold_1. +set (h' := fun a e => h e a) in *. +induction (VertexSet.elements adjsnd). simpl. assumption. +rewrite MEdgeFacts.fold_left_assoc. +set (tmp := fold_left h' l set) in *. +unfold h'. unfold h. +apply EdgeSet.remove_2. +intro. elim H0. +destruct (eq_charac _ _ H4); change_rewrite; destruct H5. +left. auto. +right. auto. assumption. + +unfold h'. unfold h. intros. +split; intros. +apply EdgeSet.remove_2. +intro H5. elim (EdgeSet.remove_1 H5 (EdgeSet.remove_3 H4)). +apply EdgeSet.remove_2. +intro H5. elim (EdgeSet.remove_1 H5 H4). +apply (EdgeSet.remove_3 (EdgeSet.remove_3 H4)). + +apply EdgeSet.remove_2. +intro H5. elim (EdgeSet.remove_1 H5 (EdgeSet.remove_3 H4)). +apply EdgeSet.remove_2. +intro H5. elim (EdgeSet.remove_1 H5 H4). +apply (EdgeSet.remove_3 (EdgeSet.remove_3 H4)). + +unfold h'. unfold h. intros. +apply RegRegProps.Equal_remove. auto. + +unfold not_incident_merge. +set (h := (fun (y : VertexSet.elt) (s' : EdgeSet.t) => + EdgeSet.remove (fst_ext e, y, Some 0%N) s')) in *. +rewrite VertexSet.fold_1. +set (h' := fun a e => h e a) in *. +generalize VertexSet.elements_2. intro. +generalize (H3 diff2). clear H3. intro HH. +induction (VertexSet.elements diff2). simpl. assumption. +rewrite MEdgeFacts.fold_left_assoc. +set (tmp := fold_left h' l s) in *. +unfold h'. unfold h. +apply EdgeSet.remove_2. +intro. +assert (VertexSet.In a diff2). +apply HH. left. intuition. +unfold diff2 in H4. +elim (VertexSet.diff_2 H4). clear H4. +apply (proj1 (edgemap_to_edgeset_charac _ _ _ (Some N0) (sym_pmap_merge_map e g q p))). +assert (eq (fst_ext e, a, Some N0) b). auto. generalize H4. clear H3 H4. intro H3. +rewrite H3. +cut (eq b (fst_ext b, snd_ext b, Some N0)). intro. +rewrite H4. +cut (In_graph_edge (fst_ext b, snd_ext b, Some N0) (merge e g p q)). intro. +destruct H5. assumption. +generalize (IE_weights _ _ H5). simpl. congruence. +rewrite <-H4. +cut (eq b (redirect (snd_ext e) (fst_ext e) b)). intro. +rewrite H5. rewrite (edge_eq _). +cut (get_weight (redirect (snd_ext e) (fst_ext e) b) = Some N0). intro. +rewrite H6. +assert (exists w, In_graph_edge + (fst_ext (redirect (snd_ext e) (fst_ext e) b), + snd_ext (redirect (snd_ext e) (fst_ext e) b), Some w) (merge e g p q)). +apply merge_5. + +rewrite H in H1. intuition. +intro. +elim H0. +destruct (eq_charac _ _ H7);change_rewrite; destruct H8. +right. auto. +left. auto. +unfold aff_edge. exists N0. rewrite <-H3. auto. + +cut (eq (fst_ext (redirect (snd_ext e) (fst_ext e) b), + snd_ext (redirect (snd_ext e) (fst_ext e) b), + None) (fst_ext b, snd_ext b, None)). intro. +unfold Interfere. rewrite H7. assumption. + +unfold redirect. destruct (OTFacts.eq_dec (fst_ext b) (snd_ext e)). +elim H0. left. auto. +destruct (OTFacts.eq_dec (snd_ext b) (snd_ext e)). +elim H0. right. auto. +apply eq_refl. + +destruct H7. +destruct H7. generalize (AE_weights _ _ H7). simpl. intro. +rewrite H8 in H7. left. auto. +generalize (IE_weights _ _ H7). simpl. congruence. + +rewrite <-H5. rewrite H4. auto. + +unfold redirect. destruct (OTFacts.eq_dec (fst_ext b) (snd_ext e)). +elim H0. left. auto. +destruct (OTFacts.eq_dec (snd_ext b) (snd_ext e)). +elim H0. right. auto. +apply eq_refl. + +apply eq_ordered_eq. constructor; simpl; try split. +apply Regs.eq_refl. +apply Regs.eq_refl. +fold (get_weight b). rewrite <-H3. simpl. apply OptionN_as_OT.eq_refl. + +apply IHl. intros. apply HH. right. auto. + +unfold h'. unfold h. intros. +split; intros. +apply EdgeSet.remove_2. +intro H5. elim (EdgeSet.remove_1 H5 (EdgeSet.remove_3 H3)). +apply EdgeSet.remove_2. +intro H5. elim (EdgeSet.remove_1 H5 H3). +apply (EdgeSet.remove_3 (EdgeSet.remove_3 H3)). + +apply EdgeSet.remove_2. +intro H5. elim (EdgeSet.remove_1 H5 (EdgeSet.remove_3 H3)). +apply EdgeSet.remove_2. +intro H5. elim (EdgeSet.remove_1 H5 H3). +apply (EdgeSet.remove_3 (EdgeSet.remove_3 H3)). + +unfold h'. unfold h. intros. +apply RegRegProps.Equal_remove. auto. + +(* last part !!!!!!! *) + +unfold AE_merge_up. +set (f := (fun (y : VertexSet.elt) (s'' : EdgeSet.t) => + EdgeSet.add (fst_ext e, y, Some 0%N) s'')) in *. +set (diff1 := (VertexSet.diff (preference_adj (fst_ext e) (merge e g p q)) + (preference_adj (fst_ext e) g))) in *. +set (diff2 := (VertexSet.diff (preference_adj (fst_ext e) g) + (preference_adj (fst_ext e) (merge e g p q)))) in *. + +set (pref := preference_adj (fst_ext e) (merge e g p q)) in *. +rewrite VertexSet.fold_1 in H1. +generalize VertexSet.elements_2. intro HH. +generalize (HH pref). clear HH. intro HH. +set (f' := fun a e => f e a) in *. induction (VertexSet.elements pref). +simpl in H1. elim (EdgeSet.empty_1 H1). + +set (set := (not_incident_merge (snd_ext e) + (not_incident_merge (fst_ext e) s diff2) + (preference_adj (snd_ext e) g))) in *. +rewrite MEdgeFacts.fold_left_assoc in H1. +set (tmp := fold_left f' l EdgeSet.empty) in *. +unfold f' in H1. unfold f in H1. +destruct (proj1 (RegRegProps.Dec.F.add_iff _ _ _) H1). +fold (eq (fst_ext e, a, Some N0) b) in H2. +rewrite <-H2. +destruct (Props.In_dec a (preference_adj (fst_ext e) g)). +cut (EdgeSet.In b set). intro Hcut. +rewrite VertexSet.fold_1. fold f'. +assert (~InA Vertex.eq a (VertexSet.elements diff1)). +intro. generalize (VertexSet.elements_2 H3). intro H4. +unfold diff1 in H4. elim (VertexSet.diff_2 H4 i). +induction (VertexSet.elements diff1). simpl. +rewrite H2. assumption. +rewrite MEdgeFacts.fold_left_assoc. +set (tmp' := fold_left f' l0 set) in *. +unfold f'. unfold f. +destruct (Vertex.eq_dec a a0). +apply EdgeSet.add_1. apply eq_ordered_eq. constructor; simpl. split; intuition. +apply OptionN_as_OT.eq_refl. +apply EdgeSet.add_2. +apply IHl0. intro. elim H3. right. auto. + +unfold f'. unfold f. intros. +apply RegRegProps.add_add. + +unfold f'. unfold f. intros. +apply RegRegProps.Dec.F.add_m. apply eq_refl. auto. + +unfold set. cut (EdgeSet.In b (not_incident_merge (fst_ext e) s diff2)). intro. +set (set' := (not_incident_merge (fst_ext e) s diff2)) in *. +unfold not_incident_merge. +set (h := (fun (y : VertexSet.elt) (s' : EdgeSet.t) => + EdgeSet.remove (snd_ext e, y, Some 0%N) s')) in *. +set (adjsnd := preference_adj (snd_ext e) g) in *. +rewrite VertexSet.fold_1. +set (h' := fun a e => h e a) in *. +induction (VertexSet.elements adjsnd). simpl. assumption. +rewrite MEdgeFacts.fold_left_assoc. +set (tmp' := fold_left h' l set') in *. +unfold h'. unfold h. +apply EdgeSet.remove_2. +intro. elim H0. +destruct (eq_charac _ _ H4); change_rewrite; destruct H5. +left. auto. +right. auto. assumption. + +unfold h'. unfold h. intros. +split; intros. +apply EdgeSet.remove_2. +intro H5. elim (EdgeSet.remove_1 H5 (EdgeSet.remove_3 H4)). +apply EdgeSet.remove_2. +intro H5. elim (EdgeSet.remove_1 H5 H4). +apply (EdgeSet.remove_3 (EdgeSet.remove_3 H4)). + +apply EdgeSet.remove_2. +intro H5. elim (EdgeSet.remove_1 H5 (EdgeSet.remove_3 H4)). +apply EdgeSet.remove_2. +intro H5. elim (EdgeSet.remove_1 H5 H4). +apply (EdgeSet.remove_3 (EdgeSet.remove_3 H4)). + +unfold h'. unfold h. intros. +apply RegRegProps.Equal_remove. auto. + +unfold not_incident_merge. +set (h := (fun (y : VertexSet.elt) (s' : EdgeSet.t) => + EdgeSet.remove (fst_ext e, y, Some 0%N) s')) in *. +rewrite VertexSet.fold_1. +set (h' := fun a e => h e a) in *. +generalize VertexSet.elements_2. intro. +generalize (H3 diff2). clear H3. intro HHH. +induction (VertexSet.elements diff2). simpl. +rewrite H. split. +unfold aff_edge. exists N0. rewrite <-H2. auto. +left. rewrite <-H2. +apply (proj2 (edgemap_to_edgeset_charac _ _ _ (Some N0) (sym_pmap g))). assumption. +rewrite MEdgeFacts.fold_left_assoc. +set (tmp' := fold_left h' l0 s) in *. +unfold h'. unfold h. +apply EdgeSet.remove_2. +intro. +assert (Vertex.eq a a0). +destruct (Vertex.eq_dec a a0). intuition. +assert (eq (fst_ext e, a0, Some N0) b) by auto. generalize H4. clear H3 H4. intro H3. +rewrite <-H3 in H2. destruct (eq_charac _ _ H2); change_rewrite; destruct H4. +auto. +apply Vertex.eq_trans with (y := fst_ext e); auto. +assert (~VertexSet.In a diff2). +intro. unfold diff2 in H5. elim (VertexSet.diff_2 H5). +apply HH. left. intuition. +elim H5. +rewrite H4. apply HHH. left. intuition. +apply IHl0. +intros. apply HHH. right. auto. + +unfold h'. unfold h. intros. +split; intros. +apply EdgeSet.remove_2. +intro H5. elim (EdgeSet.remove_1 H5 (EdgeSet.remove_3 H3)). +apply EdgeSet.remove_2. +intro H5. elim (EdgeSet.remove_1 H5 H3). +apply (EdgeSet.remove_3 (EdgeSet.remove_3 H3)). + +apply EdgeSet.remove_2. +intro H5. elim (EdgeSet.remove_1 H5 (EdgeSet.remove_3 H3)). +apply EdgeSet.remove_2. +intro H5. elim (EdgeSet.remove_1 H5 H3). +apply (EdgeSet.remove_3 (EdgeSet.remove_3 H3)). + +unfold h'. unfold h. intros. +apply RegRegProps.Equal_remove. auto. + +rewrite VertexSet.fold_1. +fold f'. +generalize VertexSet.elements_1. intro HHH. +generalize (HHH diff1 a). clear HHH. intro HHH. +assert (VertexSet.In a diff1). apply VertexSet.diff_3. apply HH. left. intuition. +assumption. +induction (VertexSet.elements diff1). simpl. +generalize (HHH H3). intro H4. inversion H4. +rewrite MEdgeFacts.fold_left_assoc. +set (tmp' := fold_left f' l0 set) in *. +unfold f'. unfold f. +destruct (Vertex.eq_dec a a0). +apply EdgeSet.add_1. +apply eq_ordered_eq. constructor; simpl; try split; intuition. +apply OptionN_as_OT.eq_refl. +apply EdgeSet.add_2. apply IHl0. intro. +generalize (HHH H4). intro H5. inversion H5; subst. +elim (n0 H7). +assumption. + +unfold f'. unfold f. intros. +apply RegRegProps.add_add. + +unfold f'. unfold f. intros. +apply RegRegProps.Dec.F.add_m. apply eq_refl. auto. + +apply IHl. +assumption. +intros. apply HH. right. auto. + +unfold f'. unfold f. intros. +apply RegRegProps.add_add. + +unfold f'. unfold f. intros. +apply RegRegProps.Dec.F.add_m. apply eq_refl. auto. +Qed. + +Lemma AE_aux_3 : forall e g s p q, +(forall a, EdgeSet.In a s <-> aff_edge a /\ In_graph_edge a g) -> +(forall b, EdgeSet.In b (AE (merge e g p q)) + <-> + (~incident b (snd_ext e) /\ + ((EdgeSet.In b s /\ ~Interfere (fst_ext b) (snd_ext b) (merge e g p q)) \/ + EdgeSet.In b (VertexSet.fold + (fun y s' => EdgeSet.add (fst_ext e, y, Some N0) s') + (preference_adj (fst_ext e) (merge e g p q)) + EdgeSet.empty)))). + +Proof. +intros e g s p q. generalize I. intro H. +split; intros. +generalize (proj1 (In_graph_aff_edge_in_AE _ _) H1). intro H3. +destruct H3 as [H3 H4]. +assert (~incident b (snd_ext e)) as Hinc. +intro. destruct H2; elim (merge_1 e g p q); auto. +unfold In_graph. rewrite H2. apply (proj1 (In_graph_edge_in_ext _ _ H4)). +unfold In_graph. rewrite H2. apply (proj2 (In_graph_edge_in_ext _ _ H4)). + +split. +assumption. + +assert (get_weight b = Some N0) as Hw. +apply AE_weights with (g := merge e g p q). assumption. + +clear H1. +generalize (merge_4 e _ g p q H4). intro H1. +destruct H1. destruct H1 as [H1 H2]. destruct H2 as [H2 HH2]. +assert (aff_edge x) as Haffb. +unfold same_type in HH2. destruct HH2; destruct H5;[auto|congruence]. + +unfold redirect in H2. +destruct (OTFacts.eq_dec (fst_ext x) (snd_ext e)). +right. +unfold weak_eq in H2. change_rewrite. destruct H2; destruct H2. + +set (f := (fun (y : VertexSet.elt) (s' : EdgeSet.t) => + EdgeSet.add (fst_ext e, y, Some 0%N) s')) in *. +set (set := preference_adj (fst_ext e) (merge e g p q)) in *. +rewrite VertexSet.fold_1. set (f' := fun a e => f e a) in *. +generalize VertexSet.elements_1. intro HH. +generalize (HH set (snd_ext x)). clear HH. intro HH. +induction (VertexSet.elements set). simpl. +assert (InA Vertex.eq (snd_ext x) nil). +apply HH. +unfold set. rewrite <-H5. rewrite (compat_preference_adj _ _ _ (Vertex.eq_sym H2)). +apply (proj1 (edgemap_to_edgeset_charac _ _ _ (Some N0) (sym_pmap_merge_map e g q p))). +cut (eq b (fst_ext b, snd_ext b, Some N0)). intro Heq. +rewrite <-Heq. destruct H4. assumption. +generalize (IE_weights _ _ H4). congruence. +rewrite (edge_eq b); change_rewrite. +assert (get_weight b = Some N0). +apply AE_weights with (g:= merge e g p q). rewrite In_graph_aff_edge_in_AE. +split; auto. +rewrite H6. apply eq_refl. +inversion H6. + +rewrite MEdgeFacts.fold_left_assoc. +set (tmp := fold_left f' l EdgeSet.empty) in *. +unfold f'. unfold f. +destruct (Vertex.eq_dec (snd_ext x) a). +apply EdgeSet.add_1. +apply eq_ordered_eq. +rewrite (edge_eq b). +constructor; simpl. split; intuition. +rewrite <-e0. intuition. +assert (get_weight b = Some N0). +apply AE_weights with (g:= merge e g p q). rewrite In_graph_aff_edge_in_AE. +split; intuition. +rewrite H6. apply OptionN_as_OT.eq_refl. +apply EdgeSet.add_2. +apply IHl. intro. generalize (HH H6). intro H7. +inversion H7; subst. +elim (n H9). +auto. + +unfold f'. unfold f. intros. +apply RegRegProps.add_add. + +unfold f'. unfold f. intros. +apply RegRegProps.Dec.F.add_m. apply eq_refl. auto. + +set (f := (fun (y : VertexSet.elt) (s' : EdgeSet.t) => + EdgeSet.add (fst_ext e, y, Some 0%N) s')) in *. +set (set := preference_adj (fst_ext e) (merge e g p q)) in *. +rewrite VertexSet.fold_1. set (f' := fun a e => f e a) in *. +generalize VertexSet.elements_1. intro HH. +generalize (HH set (snd_ext x)). clear HH. intro HH. +induction (VertexSet.elements set). simpl. +assert (InA Vertex.eq (snd_ext x) nil). +apply HH. +unfold set. rewrite <-H2. rewrite (compat_preference_adj _ _ _ (Vertex.eq_sym H5)). +apply (proj1 (edgemap_to_edgeset_charac _ _ _ (Some N0) (sym_pmap_merge_map e g q p))). +cut (eq b (snd_ext b, fst_ext b, Some N0)). intro Heq. +rewrite <-Heq. destruct H4. assumption. +generalize (IE_weights _ _ H4). congruence. +rewrite (edge_eq b); change_rewrite. +assert (get_weight b = Some N0). +apply AE_weights with (g:= merge e g p q). rewrite In_graph_aff_edge_in_AE. +split; auto. +rewrite H6. apply edge_comm. +inversion H6. + +rewrite MEdgeFacts.fold_left_assoc. +set (tmp := fold_left f' l EdgeSet.empty) in *. +unfold f'. unfold f. +destruct (Vertex.eq_dec (snd_ext x) a). +apply EdgeSet.add_1. +fold (eq (fst_ext e, a, Some N0) b). +rewrite edge_comm. apply eq_ordered_eq. +rewrite (edge_eq b). +constructor; simpl. split; intuition. +rewrite <-e0. intuition. +assert (get_weight b = Some N0). +apply AE_weights with (g:= merge e g p q). rewrite In_graph_aff_edge_in_AE. +split; auto. +rewrite H6. apply OptionN_as_OT.eq_refl. +apply EdgeSet.add_2. +apply IHl. intro. generalize (HH H6). intro H7. +inversion H7; subst. +elim (n H9). +auto. + +unfold f'. unfold f. intros. +apply RegRegProps.add_add. + +unfold f'. unfold f. intros. +apply RegRegProps.Dec.F.add_m. apply eq_refl. auto. + +destruct (OTFacts.eq_dec (snd_ext x) (snd_ext e)). + +right. +unfold weak_eq in H2. change_rewrite. destruct H2; destruct H2. +assert (get_weight x = Some N0). +apply AE_weights with (g:=g). apply (proj2 (In_graph_aff_edge_in_AE _ _)). +split; auto. + +set (f := (fun (y : VertexSet.elt) (s' : EdgeSet.t) => + EdgeSet.add (fst_ext e, y, Some 0%N) s')) in *. +set (set := preference_adj (fst_ext e) (merge e g p q)) in *. +rewrite VertexSet.fold_1. set (f' := fun a e => f e a) in *. +generalize VertexSet.elements_1. intro HH. +generalize (HH set (fst_ext x)). clear HH. intro HH. +induction (VertexSet.elements set). simpl. +assert (InA Vertex.eq (fst_ext x) nil). +apply HH. +unfold set. rewrite <-H2. rewrite (compat_preference_adj _ _ _ (Vertex.eq_sym H5)). +apply (proj1 (edgemap_to_edgeset_charac _ _ _ (Some N0) (sym_pmap_merge_map e g q p))). +cut (eq b (snd_ext b, fst_ext b, Some N0)). intro Heq. +rewrite <-Heq. destruct H4. assumption. +generalize (IE_weights _ _ H4). congruence. +rewrite (edge_eq b); change_rewrite. +assert (get_weight b = Some N0). +apply AE_weights with (g:= merge e g p q). rewrite In_graph_aff_edge_in_AE. +split; auto. +rewrite H7. apply edge_comm. +inversion H7. + +rewrite MEdgeFacts.fold_left_assoc. +set (tmp := fold_left f' l EdgeSet.empty) in *. +unfold f'. unfold f. +destruct (Vertex.eq_dec (fst_ext x) a). +apply EdgeSet.add_1. +fold (eq (fst_ext e, a, Some N0) b). +rewrite edge_comm. apply eq_ordered_eq. +rewrite (edge_eq b). +constructor; simpl. split; intuition. +rewrite <-e0. intuition. +rewrite Hw. auto. apply OptionN_as_OT.eq_refl. +apply EdgeSet.add_2. +apply IHl. intro. generalize (HH H7). intro H8. +inversion H8; subst. +elim (n0 H10). +auto. + +unfold f'. unfold f. intros. +apply RegRegProps.add_add. + +unfold f'. unfold f. intros. +apply RegRegProps.Dec.F.add_m. apply eq_refl. auto. + +set (f := (fun (y : VertexSet.elt) (s' : EdgeSet.t) => + EdgeSet.add (fst_ext e, y, Some 0%N) s')) in *. +set (set := preference_adj (fst_ext e) (merge e g p q)) in *. +rewrite VertexSet.fold_1. set (f' := fun a e => f e a) in *. +generalize VertexSet.elements_1. intro HH. +generalize (HH set (fst_ext x)). clear HH. intro HH. +induction (VertexSet.elements set). simpl. +assert (InA Vertex.eq (fst_ext x) nil). +apply HH. +unfold set. rewrite <-H5. rewrite (compat_preference_adj _ _ _ (Vertex.eq_sym H2)). +apply (proj1 (edgemap_to_edgeset_charac _ _ _ (Some N0) (sym_pmap_merge_map e g q p))). +cut (eq b (fst_ext b, snd_ext b, Some N0)). intro Heq. +rewrite <-Heq. destruct H4. assumption. +generalize (IE_weights _ _ H4). congruence. +rewrite (edge_eq b); change_rewrite. +assert (get_weight b = Some N0). +apply AE_weights with (g:= merge e g p q). rewrite In_graph_aff_edge_in_AE. +split; auto. +rewrite H6. apply eq_refl. +inversion H6. + +rewrite MEdgeFacts.fold_left_assoc. +set (tmp := fold_left f' l EdgeSet.empty) in *. +unfold f'. unfold f. +destruct (Vertex.eq_dec (fst_ext x) a). +apply EdgeSet.add_1. +apply eq_ordered_eq. +rewrite (edge_eq b). +constructor; simpl. split; intuition. +rewrite H5. intuition. +rewrite Hw. apply OptionN_as_OT.eq_refl. +apply EdgeSet.add_2. +apply IHl. intro. generalize (HH H6). intro H7. +inversion H7; subst. +elim (n0 H9). +auto. + +unfold f'. unfold f. intros. +apply RegRegProps.add_add. + +unfold f'. unfold f. intros. +apply RegRegProps.Dec.F.add_m. apply eq_refl. auto. + +assert (eq b x) as Heq. +unfold weak_eq in H2. destruct H2; destruct H2. +apply eq_ordered_eq. constructor; simpl. +inversion H2; auto. +fold (get_weight b). fold (get_weight x). +rewrite Hw. +assert (get_weight x = Some N0). +apply AE_weights with (g:=g). +rewrite In_graph_aff_edge_in_AE. +split; auto. +rewrite H6. apply OptionN_as_OT.eq_refl. +rewrite (edge_eq b). rewrite (edge_eq x). +rewrite edge_comm. apply eq_ordered_eq. constructor; simpl. +inversion H2; simpl in *; intuition. +rewrite H6. rewrite H7. auto. +rewrite H6. rewrite H7. auto. +assert (get_weight x = Some N0). +apply AE_weights with (g:=g). +rewrite In_graph_aff_edge_in_AE. +split; intuition. +rewrite H6. rewrite Hw. apply OptionN_as_OT.eq_refl. +left. split. +rewrite Heq. +rewrite (H0 x). split. +exists N0. rewrite <-Heq. auto. auto. +intro. elim (simple_graph (merge e g p q) (fst_ext b) (snd_ext b)). +split. +unfold Interfere in H5. +apply (proj1 (edgemap_to_edgeset_charac _ _ _ None (sym_imap_merge_map e g q p))). +rewrite edge_comm. +destruct H5. +generalize (AE_weights _ _ H5). simpl. congruence. +assumption. +apply (proj1 (edgemap_to_edgeset_charac _ _ _ (Some N0) (sym_pmap_merge_map e g q p))). +rewrite edge_comm. +cut (eq (fst_ext b, snd_ext b, Some N0) b). intro Heq2. +rewrite Heq2. +destruct H4. +assumption. +generalize (IE_weights _ _ H4). congruence. +rewrite (edge_eq b). change_rewrite. rewrite Hw. apply eq_refl. + +(* second part !!! *) + +destruct H1. +cut (get_weight b = Some N0). intro Hw. +assert (eq b (fst_ext b, snd_ext b, Some N0)). +rewrite (edge_eq b); change_rewrite. +apply eq_ordered_eq. constructor; simpl. +split; intuition. +rewrite Hw. apply OptionN_as_OT.eq_refl. + +destruct H2. +rewrite H3. +rewrite In_graph_aff_edge_in_AE. split. +exists N0; simpl; auto. +cut (eq b (redirect (snd_ext e) (fst_ext e) b)). intro Hcut. +rewrite <-H3. rewrite Hcut. rewrite (edge_eq _). +cut (get_weight (redirect (snd_ext e) (fst_ext e) b) = Some N0). intro Hw'. +cut (Prefere (fst_ext (redirect (snd_ext e) (fst_ext e) b)) + (snd_ext (redirect (snd_ext e) (fst_ext e) b)) + (merge e g p q)). intro. +unfold Prefere in H4. destruct H4. +assert (get_weight ((fst_ext (redirect (snd_ext e) (fst_ext e) b), + snd_ext (redirect (snd_ext e) (fst_ext e) b), Some x)) = Some N0). +apply (AE_weights) with (g:=merge e g p q). +rewrite In_graph_aff_edge_in_AE. split. +exists x. auto. auto. +rewrite Hw'. rewrite <-H5. simpl. auto. +destruct H2. apply merge_5. +apply (proj1 (H0 _) H2). +intro. elim H1. +destruct (eq_charac _ _ H5). right; intuition. left; intuition. +exists N0; simpl; auto. +intro. unfold Interfere in H5. +cut (eq (fst_ext (redirect (snd_ext e) (fst_ext e) b), + snd_ext (redirect (snd_ext e) (fst_ext e) b), + None) + (fst_ext b, snd_ext b, None)). intro H6. +rewrite H6 in H5. +elim H4. unfold Interfere. assumption. +unfold redirect; change_rewrite. +destruct (OTFacts.eq_dec (fst_ext b) (snd_ext e)). +elim H1. left. auto. +destruct (OTFacts.eq_dec (snd_ext b) (snd_ext e)). +elim H1. right. auto. +apply eq_refl. +rewrite <-Hcut. assumption. +unfold redirect; change_rewrite. +destruct (OTFacts.eq_dec (fst_ext b) (snd_ext e)). +elim H1. left. auto. +destruct (OTFacts.eq_dec (snd_ext b) (snd_ext e)). +elim H1. right. auto. +apply eq_refl. + +rewrite H3. +apply (proj2 (edgemap_to_edgeset_charac _ _ _ (Some N0) (sym_pmap_merge_map e g q p))). +set (f := (fun (y : VertexSet.elt) (s' : EdgeSet.t) => + EdgeSet.add (fst_ext e, y, Some 0%N) s')) in *. +set (s' := preference_adj (fst_ext e) (merge e g p q)) in *. +rewrite VertexSet.fold_1 in H2. +generalize VertexSet.elements_2. intro HH. +generalize (HH s'). clear HH. intro HH. +induction (VertexSet.elements s'). simpl in H2. +elim (EdgeSet.empty_1 H2). +set (f' := fun a e => f e a) in *. +rewrite MEdgeFacts.fold_left_assoc in H2. +set (tmp := fold_left f' l EdgeSet.empty) in *. +unfold f' in H2. unfold f in H2. +destruct (proj1 (RegRegProps.Dec.F.add_iff _ _ _) H2). +fold (eq (fst_ext e, a, Some N0) b) in H4. +rewrite (edge_eq b) in H4. +destruct (eq_charac _ _ H4); change_rewrite. +destruct H5. +assert (VertexSet.In (snd_ext b) s'). +apply HH. left. intuition. +unfold s' in H7. +unfold adj_set. rewrite (MapFacts.find_o _ (Vertex.eq_sym H5)). +assumption. +destruct H5. +assert (VertexSet.In (fst_ext b) s'). +apply HH. left. intuition. +unfold s' in H7. +unfold preference_adj in H7. unfold adj_set in H7. rewrite (MapFacts.find_o _ H5) in H7. +apply (sym_pmap_merge_map e g q p). assumption. +apply IHl. +assumption. +intros. apply HH. right. auto. + +unfold f'. unfold f. intros. +apply RegRegProps.add_add. + +unfold f'. unfold f. intros. +apply RegRegProps.Dec.F.add_m. apply eq_refl. auto. + +destruct H2. +rewrite AE_weights with (g:=g). auto. +rewrite (In_graph_aff_edge_in_AE). rewrite <-H0. intuition. +set (f := (fun (y : VertexSet.elt) (s' : EdgeSet.t) => + EdgeSet.add (fst_ext e, y, Some 0%N) s')) in *. +set (s' := (preference_adj (fst_ext e) (merge e g p q))) in *. +rewrite VertexSet.fold_1 in H2. +set (f' := fun a e => f e a) in *. +induction (VertexSet.elements s'). simpl in H2. +elim (EdgeSet.empty_1 H2). +rewrite MEdgeFacts.fold_left_assoc in H2. +set (tmp := fold_left f' l EdgeSet.empty) in *. +unfold f' in H2. unfold f in H2. +destruct (proj1 (RegRegProps.Dec.F.add_iff _ _ _) H2). +fold (eq (fst_ext e, a, Some N0) b) in H3. rewrite <-H3. +auto. +apply IHl. assumption. + +unfold f'. unfold f. intros. +apply RegRegProps.add_add. + +unfold f'. unfold f. intros. +apply RegRegProps.Dec.F.add_m. apply eq_refl. auto. +Qed. + +Lemma AE_merge_wl_aux : forall e g s p q, +(forall a, EdgeSet.In a s <-> aff_edge a /\ In_graph_edge a g) -> +(EdgeSet.Equal (AE_merge_up (preference_adj (fst_ext e) (merge e g p q)) + (preference_adj (fst_ext e) g) + (preference_adj (snd_ext e) g) + e s) + (AE (merge e g p q))). + +Proof. +intros. +unfold EdgeSet.Equal. intro. rewrite (AE_aux_2 e g s p q H). + rewrite (AE_aux_3 e g s p q H). +reflexivity. +Qed. + +Lemma AE_merge_wl : forall e g s p q, +(forall a, EdgeSet.In a s <-> aff_edge a /\ In_graph_edge a g) -> +(forall b, EdgeSet.In b (AE_merge_up (preference_adj (fst_ext e) (merge e g p q)) + (preference_adj (fst_ext e) g) + (preference_adj (snd_ext e) g) + e s) + <-> + aff_edge b /\ In_graph_edge b (merge e g p q)). + +Proof. +intros. +rewrite <-(In_graph_aff_edge_in_AE). +apply AE_merge_wl_aux; assumption. +Qed. + +(* It implements the interface *) + +(* Specification of the interference neighborhood *) +Lemma in_interf : forall x y g, +VertexSet.In x (interference_adj y g) <-> In_graph_edge (x,y,None) g. + +Proof. +split; intros. +apply in_interf_interf. auto. +apply interf_in_interf. auto. +Qed. + +(* Specification of the preference neighborhood *) +Lemma in_pref : forall x y g, +VertexSet.In x (preference_adj y g) <-> exists w, In_graph_edge (x,y,Some w) g. + +Proof. +split; intros. +apply in_pref_pref. auto. +destruct H. apply pref_in_pref with (w:=x0). auto. +Qed. \ No newline at end of file diff --git a/backend/InterfGraphProperties.v b/backend/InterfGraphProperties.v new file mode 100755 index 00000000..09e1b59f --- /dev/null +++ b/backend/InterfGraphProperties.v @@ -0,0 +1,8 @@ +Require Import InterfGraph. +Require Import Coloring. + +Lemma set_reg_reg_diff_ext : forall x f live live0, +SetRegReg.In x (interf_reg_reg (interf_graph f live live0)) -> fst x <> snd x. + +Proof. +Admitted. \ No newline at end of file diff --git a/backend/InterfGraph_Construction.v b/backend/InterfGraph_Construction.v new file mode 100755 index 00000000..083a80be --- /dev/null +++ b/backend/InterfGraph_Construction.v @@ -0,0 +1,189 @@ +Require Import Coqlib. +Require Import FSets. +Require Import FSetAVL. +Require Import Maps. +Require Import Ordered. +Require Import Registers. +Require Import Locations. +Require Import AST. +Require Import Op. +Require Import RTLtyping. +Require Import RTL. +Require Import Conventions. +Require Import InterfGraph. + +Definition add_interf_live + (filter: reg -> bool) (res: reg) (live: Regset.t) (g: graph): graph := + Regset.fold + (fun r g => if filter r then add_interf r res g else g) live g. + +Definition add_interf_op + (res: reg) (live: Regset.t) (g: graph): graph := + add_interf_live + (fun r => if Reg.eq r res then false else true) + res live g. + +Definition add_interf_move + (arg res: reg) (live: Regset.t) (g: graph): graph := + add_interf_live + (fun r => + if Reg.eq r res then false else + if Reg.eq r arg then false else true) + res live g. + +Definition add_interf_destroyed + (live: Regset.t) (destroyed: list mreg) (g: graph): graph := + List.fold_left + (fun g mr => Regset.fold (fun r g => add_interf_mreg r mr g) live g) + destroyed g. + +Definition add_interfs_indirect_call + (rfun: reg) (locs: list loc) (g: graph): graph := + List.fold_left + (fun g loc => + match loc with R mr => add_interf_mreg rfun mr g | _ => g end) + locs g. + +Definition add_interf_call + (ros: reg + ident) (locs: list loc) (g: graph): graph := + match ros with + | inl rfun => add_interfs_indirect_call rfun locs g + | inr idfun => g + end. + +Fixpoint add_prefs_call + (args: list reg) (locs: list loc) (g: graph) {struct args} : graph := + match args, locs with + | a1 :: al, l1 :: ll => + add_prefs_call al ll + (match l1 with R mr => add_pref_mreg a1 mr g | _ => g end) + | _, _ => g + end. + +Definition add_interf_entry + (params: list reg) (live: Regset.t) (g: graph): graph := + List.fold_left (fun g r => add_interf_op r live g) params g. + +Fixpoint add_interf_params + (params: list reg) (g: graph) {struct params}: graph := + match params with + | nil => g + | p1 :: pl => + add_interf_params pl + (List.fold_left + (fun g r => if Reg.eq r p1 then g else add_interf r p1 g) + pl g) + end. + +Definition add_edges_instr + (sig: signature) (i: instruction) (live: Regset.t) (g: graph) : graph := + match i with + | Iop op args res s => + if Regset.mem res live then + match is_move_operation op args with + | Some arg => + add_pref arg res (add_interf_move arg res live g) + | None => + add_interf_op res live g + end + else g + | Iload chunk addr args dst s => + if Regset.mem dst live + then add_interf_op dst live g + else g + | Icall sig ros args res s => + let largs := loc_arguments sig in + let lres := loc_result sig in + add_prefs_call args largs + (add_pref_mreg res lres + (add_interf_op res live + (add_interf_call ros largs + (add_interf_destroyed + (Regset.remove res live) destroyed_at_call_regs g)))) + | Itailcall sig ros args => + let largs := loc_arguments sig in + add_prefs_call args largs + (add_interf_call ros largs g) + | Ireturn (Some r) => + add_pref_mreg r (loc_result sig) g + | _ => g + end. + +Definition add_edges_instrs (f: function) (live: PMap.t Regset.t) : graph := + PTree.fold + (fun g pc i => add_edges_instr f.(fn_sig) i live!!pc g) + f.(fn_code) + empty_graph. + +Definition interf_graph (f: function) (live: PMap.t Regset.t) (live0: Regset.t) := + add_prefs_call f.(fn_params) (loc_parameters f.(fn_sig)) + (add_interf_params f.(fn_params) + (add_interf_entry f.(fn_params) live0 + (add_edges_instrs f live))). + +(** * Graph coloring *) + +(** The actual coloring of the graph is performed by a function written + directly in Caml, and not proved correct in any way. This function + takes as argument the [RTL] function, the interference graph for + this function, an assignment of types to [RTL] pseudo-registers, + and the set of all [RTL] pseudo-registers mentioned in the + interference graph. It returns the coloring as a function from + pseudo-registers to locations. *) +(* +Parameter graph_coloring: + function -> graph -> regenv -> Regset.t -> (reg -> loc). +*) +(** To ensure that the result of [graph_coloring] is a correct coloring, + we check a posteriori its result using the following Coq functions. + Let [coloring] be the function [reg -> loc] returned by [graph_coloring]. + The three properties checked are: +- [coloring r1 <> coloring r2] if there is a conflict edge between + [r1] and [r2] in the interference graph. +- [coloring r1 <> R m2] if there is a conflict edge between pseudo-register + [r1] and machine register [m2] in the interference graph. +- For all [r] mentioned in the interference graph, + the location [coloring r] is acceptable and has the same type as [r]. +*) + +Definition check_coloring_1 (g: graph) (coloring: reg -> loc) := + SetRegReg.for_all + (fun r1r2 => + if Loc.eq (coloring (fst r1r2)) (coloring (snd r1r2)) then false else true) + g.(interf_reg_reg). + +Definition check_coloring_2 (g: graph) (coloring: reg -> loc) := + SetRegMreg.for_all + (fun r1mr2 => + if Loc.eq (coloring (fst r1mr2)) (R (snd r1mr2)) then false else true) + g.(interf_reg_mreg). + +Definition same_typ (t1 t2: typ) := + match t1, t2 with + | Tint, Tint => true + | Tfloat, Tfloat => true + | _, _ => false + end. + +Definition loc_is_acceptable (l: loc) := + match l with + | R r => + if In_dec Loc.eq l temporaries then false else true + | S (Local ofs ty) => + if zlt ofs 0 then false else true + | _ => + false + end. + +Definition check_coloring_3 (rs: Regset.t) (env: regenv) (coloring: reg -> loc) := + Regset.for_all + (fun r => + let l := coloring r in + andb (loc_is_acceptable l) (same_typ (env r) (Loc.type l))) + rs. + +Definition check_coloring + (g: graph) (env: regenv) (rs: Regset.t) (coloring: reg -> loc) := + andb (check_coloring_1 g coloring) + (andb (check_coloring_2 g coloring) + (check_coloring_3 rs env coloring)). diff --git a/backend/Interference_adjacency.v b/backend/Interference_adjacency.v new file mode 100755 index 00000000..3d5d90dd --- /dev/null +++ b/backend/Interference_adjacency.v @@ -0,0 +1,66 @@ +Require Import InterfGraphMapImp. +Require Import Graph_Facts. +Require Import FSets. +Require Import SetsFacts. +Require Import Edges. + +Import Edge RegFacts Props RegRegProps. + +(* Some properties about the interference adjacency + and the same about preference adjacency *) + +(* x is not in its own interference neighborhood *) +Lemma not_in_interf_self : forall x g, +~VertexSet.In x (interference_adj x g). + +Proof. +intros x g. rewrite in_interf. intro H. +elim (In_graph_edge_diff_ext _ _ H). auto. +Qed. + +(* x is not in its own preference neighborhood *) +Lemma not_in_pref_self : forall x g, +~VertexSet.In x (preference_adj x g). + +Proof. +intros x g. rewrite in_pref. intro H. destruct H. +elim (In_graph_edge_diff_ext _ _ H). auto. +Qed. + +(* If x is an interference neighbor of y in g + then y is an interference neighbor of x in g *) +Lemma interf_adj_comm : forall x y g, +VertexSet.In x (interference_adj y g) -> VertexSet.In y (interference_adj x g). + +Proof. +intros x y g H. rewrite in_interf. rewrite edge_comm. rewrite <-in_interf. auto. +Qed. + +(* If x is a preference neighbor of y in g + then y is a preference neighbor of x in g *) +Lemma pref_adj_comm : forall x y g, +VertexSet.In x (preference_adj y g) -> VertexSet.In y (preference_adj x g). + +Proof. +intros x y g H. +rewrite in_pref in H. destruct H. rewrite edge_comm in H. +rewrite in_pref. exists x0. assumption. +Qed. + +(* If x is an interference neighbor of any vertex of g then x is in g *) +Lemma in_interf_in : forall x r g, +VertexSet.In x (interference_adj r g) -> In_graph x g. + +Proof. +intros x r g H. rewrite in_interf in H. +apply (proj1 (In_graph_edge_in_ext _ _ H)). +Qed. + +(* If x is a preferenec neighbor of any vertex then x is in g *) +Lemma in_pref_in : forall x r g, +VertexSet.In x (preference_adj r g) -> In_graph x g. + +Proof. +intros x r g H. rewrite in_pref in H. destruct H. +apply (proj1 (In_graph_edge_in_ext _ _ H)). +Qed. diff --git a/backend/Merge_Adjacency.v b/backend/Merge_Adjacency.v new file mode 100755 index 00000000..d6d879ae --- /dev/null +++ b/backend/Merge_Adjacency.v @@ -0,0 +1,347 @@ +Require Import FSets. +Require Import InterfGraphMapImp. +Require Import Remove_Vertex_Degree. +Require Import Edges. +Require Import MyRegisters. +Require Import Interference_adjacency. +Require Import Graph_Facts. + +Module Register := Regs. + +Import Edge RegFacts Props. + +(* The following lemmas define the interference adjacency + of the first endpoint of e after the merge of e *) + +Lemma merge_interf_adj_fst_1 : forall e g H0 Haff, +VertexSet.Subset (interference_adj (fst_ext e) (merge e g H0 Haff)) + (VertexSet.union (interference_adj (fst_ext e) g) + (interference_adj (snd_ext e) g)). + +Proof. +unfold VertexSet.Subset. intros e g H Haff a H1. +rewrite in_interf in H1. generalize (In_merge_edge_inv _ _ _ H Haff H1); intro. +destruct H0 as [x H0]. destruct H0 as [H3 H4]. +destruct x as [ex wx]. destruct ex as [x1 x2]. +assert (interf_edge (x1,x2,wx)) as Hinterf. +destruct H4. unfold same_type in H2. destruct H2; destruct H2. +unfold aff_edge in H4. destruct H4. simpl in H4. congruence. +assumption. +destruct H4 as [H4 HH4]. +assert (eq (a, fst_ext e, None) (redirect (snd_ext e) (fst_ext e) (x1,x2,wx))). +apply weak_eq_interf_eq. assumption. unfold interf_edge. auto. +unfold interf_edge. rewrite redirect_weight_eq. auto. clear H4. +generalize (redirect_charac (x1,x2,wx) (snd_ext e) (fst_ext e)); intros. +destruct H2. destruct H2. destruct H4. rewrite <-H2 in H0. +apply VertexSet.union_2. rewrite in_interf. rewrite H0. assumption. +destruct H2. destruct H2. rewrite <-H2 in H0. change_rewrite. +apply VertexSet.union_3. rewrite in_interf. +destruct (eq_charac _ _ H0); destruct H5; change_rewrite. +assert (eq (a, snd_ext e, None) (x1, x2, wx)). +Eq_comm_eq. apply (Register.eq_trans _ _ _ H5 H6). +generalize (get_weight_m _ _ H0). simpl. intro. rewrite H7. apply OptionN_as_OT.eq_refl. +rewrite H7. assumption. +assert (eq (a,snd_ext e, None) (x1,x2,wx)). +Eq_comm_eq. +generalize (get_weight_m _ _ H0). simpl. intro. rewrite H7. apply OptionN_as_OT.eq_refl. +rewrite H7. assumption. +destruct H2; change_rewrite. rewrite <-H2 in H0. +destruct (eq_charac _ _ H0); destruct H5; change_rewrite. +apply VertexSet.union_3. rewrite in_interf. +assert (eq (a, snd_ext e, None) (x1,x2,wx)). +Eq_eq. +generalize (get_weight_m _ _ H0). simpl. intro. rewrite H7. apply OptionN_as_OT.eq_refl. +rewrite H7. assumption. +apply VertexSet.union_3. rewrite in_interf. +assert (eq (a, snd_ext e, None) (x1,x2,wx)). +Eq_eq. +apply (Register.eq_trans _ _ _ H5 H6). +generalize (get_weight_m _ _ H0). simpl. intro. rewrite H7. apply OptionN_as_OT.eq_refl. +rewrite H7. assumption. +Qed. + +Lemma merge_interf_adj_fst_2 : forall e g H0 Haff, +VertexSet.Subset (interference_adj (fst_ext e) g) + (interference_adj (fst_ext e) (merge e g H0 Haff)). + +Proof. +unfold VertexSet.Subset. intros. +rewrite in_interf in H. rewrite in_interf. +destruct (redirect_charac (a, fst_ext e, None) (snd_ext e) (fst_ext e)); intros. +change_rewrite. destruct H1. destruct H2. +rewrite H1. apply In_merge_interf_edge. assumption. +unfold interf_edge. auto. +change_rewrite. destruct H1. destruct H1. +elim (interf_pref_conflict (snd_ext e) (fst_ext e) g). +split. unfold Prefere. destruct Haff. exists x. +rewrite edge_comm. rewrite <-H3. rewrite (edge_eq e) in H0. assumption. +unfold Interfere. +assert (eq (snd_ext e, fst_ext e, None) (a, fst_ext e, None)) by Eq_eq. +rewrite H3. assumption. +destruct H1. rewrite H1. apply In_merge_interf_edge. assumption. +unfold interf_edge. auto. +Qed. + +Lemma merge_interf_adj_fst_3 : forall e g H0 Haff, +VertexSet.Subset (interference_adj (snd_ext e) g) + (interference_adj (fst_ext e) (merge e g H0 Haff)). + +Proof. +unfold VertexSet.Subset. intros. +rewrite in_interf in H. rewrite in_interf. +destruct (redirect_charac (a, snd_ext e, None) (snd_ext e) (fst_ext e)); intros. +change_rewrite. destruct H1. destruct H2. elim H3. auto. +change_rewrite. destruct H1. destruct H1. +elim (In_graph_edge_diff_ext _ _ H). change_rewrite. auto. +destruct H1. rewrite H1. apply In_merge_interf_edge. assumption. +unfold interf_edge. auto. +Qed. + +Lemma merge_interf_adj_fst : forall e g H0 Haff, +VertexSet.Equal (interference_adj (fst_ext e) (merge e g H0 Haff)) + (VertexSet.union (interference_adj (fst_ext e) g) + (interference_adj (snd_ext e) g)). + +Proof. +intros. split; intros. +apply (merge_interf_adj_fst_1 _ _ H0 Haff _ H). +destruct (VertexSet.union_1 H). + apply (merge_interf_adj_fst_2 _ _ H0 Haff _ H1). + apply (merge_interf_adj_fst_3 _ _ H0 Haff _ H1). +Qed. + +(* The following lemmas define the equality of the interference adjacency + of any vertex which is not an endpoint of e after the merge of e *) + +Lemma merge_interf_adj_1 : forall x e g Hin Haff, +~Register.eq x (fst_ext e) -> +~Register.eq x (snd_ext e) -> +VertexSet.Subset (VertexSet.remove (snd_ext e) (interference_adj x g)) + (interference_adj x (merge e g Hin Haff)). + +Proof. +unfold VertexSet.Subset. intros. +rewrite in_interf. rewrite Dec.F.remove_iff in H1. destruct H1. +destruct (redirect_charac (a,x,None) (snd_ext e) (fst_ext e)); change_rewrite. +destruct H3. destruct H4. rewrite H3. +apply In_merge_interf_edge. rewrite <-in_interf. assumption. +unfold interf_edge. auto. +destruct H3. destruct H3. elim H2. auto. +destruct H3. elim H0. auto. +Qed. + +Lemma merge_interf_adj_2 : forall x e g Hin Haff, +~Register.eq x (fst_ext e) -> +~Register.eq x (snd_ext e) -> +VertexSet.Subset (interference_adj x (merge e g Hin Haff)) + (VertexSet.add (fst_ext e) (interference_adj x g)). + +Proof. +unfold VertexSet.Subset. intros. +destruct (Register.eq_dec a (fst_ext e)); [apply VertexSet.add_1|apply VertexSet.add_2]; intuition. +rewrite in_interf in H1. generalize (In_merge_edge_inv _ _ _ Hin Haff H1). intro. +destruct H2. destruct H2. destruct H3. +assert (eq (a,x,None) (redirect (snd_ext e) (fst_ext e) x0)). +apply weak_eq_interf_eq. assumption. unfold interf_edge. auto. +unfold interf_edge. rewrite redirect_weight_eq. +unfold same_type in H4. destruct H4; destruct H4. +unfold aff_edge in H5. destruct H5. simpl in H5. congruence. +auto. intuition. +rewrite in_interf. rewrite H5. +destruct (redirect_charac x0 (snd_ext e) (fst_ext e)). +destruct H6. destruct H7. rewrite <-H6. assumption. +destruct H6. destruct H6. rewrite <-H6 in H5. +destruct (eq_charac _ _ H5); destruct H8; change_rewrite. +elim n. auto. elim H. auto. +destruct H6. rewrite <-H6 in H5. +destruct (eq_charac _ _ H5); destruct H8; change_rewrite. +elim H. auto. elim n. auto. +Qed. + +Lemma merge_interf_adj_not_in : forall x e g Hin Haff, +~Register.eq x (fst_ext e) -> +~Register.eq x (snd_ext e) -> +~VertexSet.In x (interference_adj (snd_ext e) g) -> +VertexSet.Equal (interference_adj x (merge e g Hin Haff)) + (interference_adj x g). + +Proof. +intros. split; intros. +generalize (merge_interf_adj_2 x e g Hin Haff H H0 a H2). intro. +rewrite Dec.F.add_iff in H3. destruct H3. +rewrite <-H3 in H2. generalize (interf_adj_comm _ _ _ H2). intro. +rewrite merge_interf_adj_fst in H4. destruct (VertexSet.union_1 H4). +rewrite <-H3. apply interf_adj_comm. assumption. +elim H1. auto. +assumption. +apply merge_interf_adj_1; auto. +apply VertexSet.remove_2; auto. +intro. elim H1. apply interf_adj_comm. rewrite H3. assumption. +Qed. + +Lemma merge_interf_adj_in_snd : forall x e g Hin Haff, +~Register.eq x (fst_ext e) -> +~Register.eq x (snd_ext e) -> +VertexSet.In x (interference_adj (snd_ext e) g) -> +VertexSet.Equal (interference_adj x (merge e g Hin Haff)) + (VertexSet.add (fst_ext e) (VertexSet.remove (snd_ext e) + (interference_adj x g))). + +Proof. +intros. split; intros. +generalize (merge_interf_adj_2 _ _ _ Hin Haff H H0 _ H2). intro. +rewrite Dec.F.add_iff in H3. destruct H3. +apply VertexSet.add_1. assumption. +apply VertexSet.add_2. apply VertexSet.remove_2. +intro. rewrite <-H4 in H2. rewrite in_interf in H2. +generalize (proj1 (In_graph_edge_in_ext _ _ H2)). change_rewrite. intro. +rewrite In_merge_vertex in H5. destruct H5. elim H6. auto. assumption. + +rewrite Dec.F.add_iff in H2. destruct H2. +rewrite <-H2. apply interf_adj_comm. rewrite merge_interf_adj_fst. +apply VertexSet.union_3. auto. +apply merge_interf_adj_1; auto. +Qed. + +Lemma merge_interf_adj_in_both : forall x e g Hin Haff, +VertexSet.In x (interference_adj (snd_ext e) g) -> +VertexSet.In x (interference_adj (fst_ext e) g) -> +VertexSet.Equal (interference_adj x (merge e g Hin Haff)) + (VertexSet.remove (snd_ext e) (interference_adj x g)). + +Proof. +intros. +assert (~Register.eq x (fst_ext e)). +intro. rewrite H1 in H0. elim (not_in_interf_self _ _ H0). +assert (~Register.eq x (snd_ext e)). +intro. rewrite H2 in H. elim (not_in_interf_self _ _ H). +rewrite merge_interf_adj_in_snd; auto. +split; intros. +rewrite Dec.F.add_iff in H3. destruct H3. +apply VertexSet.remove_2. intro. +elim (In_graph_edge_diff_ext _ _ Hin). rewrite H3. auto. +rewrite <-H3. apply interf_adj_comm. assumption. +assumption. +apply VertexSet.add_2. auto. +Qed. + +Lemma preference_adj_merge : forall x e g p q, +~Register.eq x (fst_ext e) -> +~Register.eq x (snd_ext e) -> +VertexSet.Subset (preference_adj x (merge e g p q)) + (VertexSet.add (fst_ext e) + (VertexSet.remove (snd_ext e) (preference_adj x g))). + +Proof. +intros x e g p q H0 H1. generalize I. intro H. +unfold VertexSet.Subset; intros. +rewrite in_pref in H2. destruct H2. +generalize (In_merge_edge_inv e _ g p q H2). intro. +destruct H3. destruct H3. destruct H4 as [H4 HH4]. +unfold weak_eq in H4. change_rewrite. destruct H4. +unfold redirect in H4; change_rewrite. +destruct (OTFacts.eq_dec (fst_ext x1) (snd_ext e)); change_rewrite. +destruct H4. apply VertexSet.add_1. intuition. +destruct (OTFacts.eq_dec (snd_ext x1) (snd_ext e)); change_rewrite. +destruct H4. elim H0. intuition. +destruct H4. apply VertexSet.add_2. apply VertexSet.remove_2. +intro. elim n. rewrite H6. intuition. +rewrite in_pref. unfold same_type in HH4. +destruct HH4. destruct H6. +destruct H6. exists x2. +assert (eq (a,x,Some x2) x1). +apply eq_ordered_eq. split; try split; simpl; auto. +unfold get_weight in H6. rewrite H6. apply OptionN_as_OT.eq_refl. +rewrite H8. assumption. +destruct H6. unfold interf_edge in H7. simpl in H7. congruence. +unfold redirect in H4; change_rewrite. +destruct (OTFacts.eq_dec (fst_ext x1) (snd_ext e)); change_rewrite. +destruct H4. elim H0. auto. +destruct (OTFacts.eq_dec (snd_ext x1) (snd_ext e)); change_rewrite. +destruct H4. apply VertexSet.add_1. intuition. +destruct H4. apply VertexSet.add_2. apply VertexSet.remove_2. +intro. elim n0. rewrite H6. intuition. +rewrite in_pref. unfold same_type in HH4. +destruct HH4. destruct H6. +destruct H6. exists x2. +assert (eq (a,x,Some x2) x1). +rewrite edge_comm. apply eq_ordered_eq. split; try split; simpl; auto. +unfold get_weight in H6. rewrite H6. apply OptionN_as_OT.eq_refl. +rewrite H8. assumption. +destruct H6. unfold interf_edge in H7. simpl in H7. congruence. +Qed. + +Lemma preference_adj_merge_2 : forall x e g p q, +~Register.eq x (fst_ext e) -> +~Register.eq x (snd_ext e) -> +VertexSet.Subset (VertexSet.remove (fst_ext e) (VertexSet.remove (snd_ext e) (preference_adj x g))) + (preference_adj x (merge e g p q)). + +Proof. +unfold VertexSet.Subset; intros x e g p q H0 H1 a H2. generalize I. intro H. +assert (~Register.eq a (fst_ext e)). intro. +rewrite H3 in H2. elim (VertexSet.remove_1 (Register.eq_refl _) H2). +generalize (VertexSet.remove_3 H2). clear H2. intro H2. +assert (~Register.eq a (snd_ext e)). intro. +rewrite H4 in H2. elim (VertexSet.remove_1 (Register.eq_refl _) H2). +generalize (VertexSet.remove_3 H2). clear H2. intro. +rewrite in_pref in H2. destruct H2. +assert (exists w, In_graph_edge (a,x,Some w) (merge e g p q)). +cut (Prefere (fst_ext (redirect (snd_ext e) (fst_ext e) (a,x,Some x0))) + (snd_ext (redirect (snd_ext e) (fst_ext e) (a,x,Some x0))) + (merge e g p q)). intro. +unfold redirect in H5; change_rewrite. +destruct (OTFacts.eq_dec a (snd_ext e)); change_rewrite. +elim (H4 r). +destruct (OTFacts.eq_dec x (snd_ext e)); change_rewrite. +elim (H1 r). +unfold Prefere in H5. assumption. +apply In_merge_pref_edge. assumption. +intro. destruct (eq_charac _ _ H5); change_rewrite; destruct H6. +elim H1. auto. +elim H0. auto. +unfold aff_edge. exists x0. auto. +intro. unfold redirect in H5;change_rewrite. +destruct (OTFacts.eq_dec a (snd_ext e)); change_rewrite. +elim (H4 r). +destruct (OTFacts.eq_dec x (snd_ext e)); change_rewrite. +elim (H1 r). +unfold Interfere in H5. generalize (In_merge_edge_inv e _ g p q H5). intro H8. +generalize I. intro H7. +destruct H8. destruct H6. destruct H8. +unfold weak_eq in H8. change_rewrite. destruct H8. +unfold redirect in H8; change_rewrite. +destruct (OTFacts.eq_dec (fst_ext x1) (snd_ext e)); change_rewrite; destruct H8. +elim (H3 H8). +destruct (OTFacts.eq_dec (snd_ext x1) (snd_ext e)); change_rewrite. +elim (H0 H10). +elim (interf_pref_conflict a x g). +split. +unfold Prefere. exists x0. assumption. +unfold Interfere. +assert (eq (a,x,None) x1). +apply weak_eq_interf_eq. unfold weak_eq. change_rewrite. +left. split; auto. unfold interf_edge. auto. +unfold same_type in H9. destruct H9. destruct H9. +unfold aff_edge in H11. destruct H11. simpl in H11. congruence. +destruct H9. assumption. +rewrite H11. assumption. + +unfold redirect in H8; change_rewrite. +destruct (OTFacts.eq_dec (fst_ext x1) (snd_ext e)); change_rewrite. +destruct H8. elim (H0 H10). +destruct (OTFacts.eq_dec (snd_ext x1) (snd_ext e)); change_rewrite. +destruct H8. elim (H3 H8). +elim (interf_pref_conflict a x g). +split. +unfold Prefere. exists x0. assumption. +unfold Interfere. destruct H8. +assert (eq (a,x,None) x1). +apply weak_eq_interf_eq. unfold weak_eq. change_rewrite. +right. split; auto. unfold interf_edge. auto. +unfold same_type in H9. destruct H9. destruct H9. +unfold aff_edge in H11. destruct H11. simpl in H11. congruence. +destruct H9. assumption. +rewrite H11. assumption. +rewrite in_pref. assumption. +Qed. diff --git a/backend/Merge_Degree.v b/backend/Merge_Degree.v new file mode 100755 index 00000000..15849794 --- /dev/null +++ b/backend/Merge_Degree.v @@ -0,0 +1,205 @@ +Require Import FSets. +Require Import InterfGraphMapImp. +Require Import Merge_Adjacency. +Require Import ZArith. +Require Import Edges. +Require Import MyRegisters. +Require Import Merge_Adjacency. +Require Import Interference_adjacency. +Require Import Remove_Vertex_Degree. + +Module Register := Regs. + +Import Edge Props RegFacts. + +(* If x interferes with both (fst_ext e) and (snd_ext e), then + its degree decreases of one when e is coalesced *) +Lemma merge_degree_dec_inter : forall x e g Hin Haff, +VertexSet.In x (interference_adj (fst_ext e) g) -> +VertexSet.In x (interference_adj (snd_ext e) g) -> +interf_degree g x = S (interf_degree (merge e g Hin Haff) x). + +Proof. +intros. unfold interf_degree. +rewrite (cardinal_m (merge_interf_adj_in_both _ _ _ Hin Haff H0 H)). +rewrite remove_cardinal_1. reflexivity. apply interf_adj_comm. auto. +Qed. + +(* If x does not interfere with both (fst_ext e) and (snd_ext e) + then its degree is unchanged when e is coalesced *) +Lemma merge_dec_eq : forall x e g Hin Haff, +~VertexSet.In x (interference_adj (fst_ext e) g) \/ +~VertexSet.In x (interference_adj (snd_ext e) g) -> +~Register.eq x (fst_ext e) -> +~Register.eq x (snd_ext e) -> +interf_degree g x = interf_degree (merge e g Hin Haff) x. + +Proof. +intros. unfold interf_degree. destruct H. +destruct (In_dec x (interference_adj (snd_ext e) g)). +rewrite (cardinal_m (merge_interf_adj_in_snd _ _ _ Hin Haff H0 H1 i)). +rewrite add_cardinal_2. rewrite remove_cardinal_1. reflexivity. +apply interf_adj_comm. auto. +intro. elim H. apply interf_adj_comm. apply (VertexSet.remove_3 H2). +rewrite (cardinal_m (merge_interf_adj_not_in _ _ _ Hin Haff H0 H1 n)). reflexivity. +rewrite (cardinal_m (merge_interf_adj_not_in _ _ _ Hin Haff H0 H1 H)). reflexivity. +Qed. + +(* The interference degree of any vertex which is not an endpoint + of e decreases when e is coalesced *) +Lemma merge_degree_dec : forall x e g Hin Haff, +~Register.eq x (fst_ext e) -> +~Register.eq x (snd_ext e) -> +interf_degree (merge e g Hin Haff) x <= interf_degree g x. + +Proof. +intros. destruct (In_dec x (interference_adj (fst_ext e) g)). +destruct (In_dec x (interference_adj (snd_ext e) g)). +rewrite (merge_degree_dec_inter x e g Hin Haff i i0). auto. +rewrite (merge_dec_eq x e g Hin Haff (or_intror _ n)); auto. +rewrite (merge_dec_eq x e g Hin Haff (or_introl _ n)); auto. +Qed. + +(* If x does not interfere with the first endpoint of e then + x is of low-degree in (merge e g Hin Haff) iff it is in g *) +Lemma low_merge_low_fst : forall x e g K Hin Haff, +~VertexSet.In x (interference_adj (fst_ext e) g) -> +~Register.eq x (snd_ext e) -> +~Register.eq x (fst_ext e) -> +has_low_degree (merge e g Hin Haff) K x = has_low_degree g K x. + +Proof. +intros x e g palette Hin Haff H H0 H1. unfold has_low_degree. +rewrite (merge_dec_eq x e g Hin Haff); auto. +Qed. + +(* If x does not interfere with the second endpoint of e then + x is of low-degree in (merge e g Hin Haff) iff it is in g *) +Lemma low_merge_low_snd : forall x e g K Hin Haff, +~VertexSet.In x (interference_adj (snd_ext e) g) -> +~Register.eq x (snd_ext e) -> +~Register.eq x (fst_ext e) -> +has_low_degree (merge e g Hin Haff) K x = has_low_degree g K x. + +Proof. +intros x e g palette Hin Haff H H0 H1. unfold has_low_degree. +rewrite (merge_dec_eq x e g Hin Haff); auto. +Qed. + +(* A high-degree vertex of (merge e g Hin Haff) is of high-degree in g *) +Lemma merge_low_1 : forall g e x K Haff Hin, +has_low_degree (merge e g Hin Haff) K x = false -> +~Register.eq x (snd_ext e) -> +~Register.eq x (fst_ext e) -> +has_low_degree g K x = false. + +Proof. +intros g e x K Haff Hin Hpre H0 H1. unfold has_low_degree in *. +destruct (le_lt_dec K (interf_degree (merge e g Hin Haff) x)); +destruct (le_lt_dec K (interf_degree g x )); inversion Hpre. +reflexivity. +generalize (merge_degree_dec x e g Hin Haff H1 H0). intro. +apply False_ind. intuition. +Qed. + +(* A low-degree vertex of g is of low-degree in (merge e g Haff Hin) *) +Lemma low_dec : forall x e g Hin Haff K, +~Register.eq x (fst_ext e) -> +~Register.eq x (snd_ext e) -> +has_low_degree g K x = true -> +has_low_degree (merge e g Hin Haff) K x = true. + +Proof. +intros. +case_eq (has_low_degree (merge e g Hin Haff) K x);[auto|intros]. +rewrite (merge_low_1 g e x K Haff Hin H2 H0 H) in H1. inversion H1. +Qed. + +(* A vertex high-degree vertex of g (which is not an endpoint of e) + which is of low-degree in (merge e g p q) belongs to the interference + neighborhood of the two endpoints of e in g *) +Lemma merge_dec_interf : forall x e k g p q, +has_low_degree g k x = false -> +has_low_degree (merge e g p q) k x = true -> +~Register.eq x (fst_ext e) -> +~Register.eq x (snd_ext e) -> +VertexSet.In x (interference_adj (fst_ext e) g) /\ +VertexSet.In x (interference_adj (snd_ext e) g). + +Proof. +intros. +destruct (In_dec x (interference_adj (fst_ext e) g)). +split. assumption. +destruct (In_dec x (interference_adj (snd_ext e) g)). +assumption. +rewrite (low_merge_low_snd _ _ _ _ p q n H2 H1) in H0. rewrite H0 in H. inversion H. +rewrite (low_merge_low_fst _ _ _ _ p q n H2 H1) in H0. rewrite H0 in H. inversion H. +Qed. + +(* A vertex high-degree vertex of g (which is not an endpoint of e) + which is of low-degree in (merge e g p q) is of degree k in g *) +Lemma merge_dec_K : forall x e k g p q, +~Register.eq x (fst_ext e) -> +~Register.eq x (snd_ext e) -> +has_low_degree g k x = false -> +has_low_degree (merge e g p q) k x = true -> +interf_degree g x = k. + +Proof. +intros x e k g p q H0 H1 H2 H3. generalize I. intro H. unfold interf_degree. +assert (VertexSet.In x (interference_adj (fst_ext e) g) /\ + VertexSet.In x (interference_adj (snd_ext e) g)). +apply (merge_dec_interf x e k g p q); auto. +destruct H4. +generalize (merge_degree_dec_inter x e g p q H4 H5). intro. +unfold has_low_degree, interf_degree in *. +destruct (le_lt_dec k (VertexSet.cardinal (interference_adj x (merge e g p q)))). +inversion H3. +destruct (le_lt_dec k (VertexSet.cardinal (interference_adj x g))). +rewrite H6 in l0. rewrite (lt_le_S_eq _ _ l l0). assumption. +inversion H2. +Qed. + +(* Reciprocally, a vertex of degree k interfering with + the two endpoints of g is of low-degree in (merge e g p q) *) +Lemma merge_dec_low : forall x e k g p q, +interf_degree g x = k -> +VertexSet.In x (interference_adj (fst_ext e) g) -> +VertexSet.In x (interference_adj (snd_ext e) g) -> +has_low_degree (merge e g p q) k x = true. + +Proof. +unfold interf_degree. intros x e k g p q H0 H1 H2. generalize I. intro H. +assert (~Register.eq x (fst_ext e)). +intro. elim (not_in_interf_self (fst_ext e) g). rewrite H3 in H1. assumption. +assert (~Register.eq x (snd_ext e)). +intro. elim (not_in_interf_self (snd_ext e) g). rewrite H4 in H2. assumption. +generalize (merge_degree_dec_inter x e g p q H1 H2). intro. +unfold has_low_degree, interf_degree in *. +destruct (le_lt_dec k (VertexSet.cardinal (interference_adj x (merge e g p q)))). +rewrite H5 in H0. rewrite <-H0 in l. elim (le_S_irrefl _ l). +reflexivity. +Qed. + +(* Again, unused but meaningful theorem, summarizing evolution of degree + when an edge e is coalesced *) + +Theorem merge_degree_evolution : forall x e k g p q, +~Register.eq x (fst_ext e) -> +~Register.eq x (snd_ext e) -> +((has_low_degree g k x = false /\ has_low_degree (merge e g p q) k x = true) + <-> +(interf_degree g x = k /\ + VertexSet.In x (interference_adj (fst_ext e) g) /\ + VertexSet.In x (interference_adj (snd_ext e) g))). + +Proof. +split; intros. +destruct H1. +split. apply (merge_dec_K x e k g p q); auto. +apply (merge_dec_interf x e k g p q); auto. +destruct H1. destruct H2. +split. unfold has_low_degree. rewrite H1. +destruct (le_lt_dec k k). reflexivity. elim (lt_irrefl _ l). +apply merge_dec_low; auto. +Qed. diff --git a/backend/Merge_Move.v b/backend/Merge_Move.v new file mode 100755 index 00000000..1e20d001 --- /dev/null +++ b/backend/Merge_Move.v @@ -0,0 +1,730 @@ +Require Import FSetInterface. +Require Import InterfGraphMapImp. +Require Import Remove_Vertex_Move. +Require Import ZArith. +Require Import Edges. +Require Import MyRegisters. +Require Import Affinity_relation. +Require Import Interference_adjacency. +Require Import Graph_Facts. +Require Import Merge_Adjacency. + +Import Edge Props RegFacts. + +(* A nonmove-related vertex of g different from the first endpoint of e + is nonmove-related in (merge e g p q) *) +Lemma move_merge_false : forall x e g p q, +~Register.eq x (fst_ext e) -> +move_related g x = false -> +move_related (merge e g p q) x = false. + +Proof. +intros x e g p q Hfst H0. generalize I. intro H. +case_eq (move_related (merge e g p q) x); intros. +generalize (move_related_charac _ _ H1). intro. +destruct H2. destruct H2. destruct H3. +generalize (In_merge_edge_inv _ _ _ p q H3). intro. +destruct H5. destruct H5. +assert (move_related g x = true). +apply move_related_charac2 with (e := x1). +unfold weak_eq in H6. destruct H6. +unfold same_type in H7. destruct H7. destruct H7. assumption. +destruct H7. destruct H2. unfold interf_edge in H2. congruence. assumption. +destruct H6. +unfold redirect in H6. +destruct (OTFacts.eq_dec (fst_ext x1) (snd_ext e)); destruct H6; change_rewrite; destruct H6. +destruct H4. +elim Hfst. rewrite H4. auto. +right. rewrite H4. auto. +destruct (OTFacts.eq_dec (snd_ext x1) (snd_ext e)); change_rewrite. +destruct H4. +right. rewrite H4. auto. +elim Hfst. rewrite H4. auto. +destruct H4. +right. rewrite H4. auto. +elim Hfst. rewrite H4. auto. +destruct (OTFacts.eq_dec (snd_ext x1) (snd_ext e)); change_rewrite. +destruct H4. +left. rewrite H4. auto. +elim Hfst. rewrite H4. auto. +destruct H4. +left. rewrite H4. auto. +right. rewrite H4. auto. +destruct (OTFacts.eq_dec (snd_ext x1) (snd_ext e)); change_rewrite. +destruct H4. +elim Hfst. rewrite H4. auto. +left. rewrite H4. auto. +destruct H4. +right. rewrite H4. auto. +left. rewrite H4. auto. +rewrite H0 in H7. inversion H7. +auto. +Qed. + +(* Equivalently, a move-related vertex of (merge e g p q) is + move-related in g *) +Lemma move_related_merge_move_related : forall g e x H H0, +move_related (merge e g H H0) x = true -> +move_related g x = true. + +Proof. +intros. case_eq (move_related g x); auto. +intro. assert (~Register.eq x (fst_ext e)). +intro. rewrite (compat_bool_move _ _ _ H3) in H2. +rewrite (proj1 (Aff_edge_aff _ _ H H0)) in H2. inversion H2. +rewrite (move_merge_false _ _ _ H H0 H3 H2) in H1. inversion H1. +Qed. + +(* If x is neither a neighbor of (fst_ext e) nor (snd_ext e) and + if x is move-related in g, then x is move-related in (merge e g H Haff) *) +Lemma move_related_move_related_merge : forall g e x Haff H, +~VertexSet.In x (preference_adj (fst_ext e) g) -> +~VertexSet.In x (preference_adj (snd_ext e) g) -> +move_related g x = true -> +move_related (merge e g H Haff) x = true. + +Proof. +intros. +generalize (move_related_charac _ _ H2). intro. +destruct H3. destruct H3. destruct H4. +assert (~Register.eq x (fst_ext e)). +intro. elim H1. rewrite H6. +rewrite in_pref. destruct Haff. exists x1. rewrite <-H7. +rewrite (edge_eq e) in H. assumption. +assert (~Register.eq x (snd_ext e)). +intro. elim H0. rewrite H7. +rewrite in_pref. destruct Haff. exists x1. rewrite <-H8. +rewrite (edge_eq e) in H. rewrite edge_comm. assumption. + +assert (Prefere (fst_ext (redirect (snd_ext e) (fst_ext e) x0)) + (snd_ext (redirect (snd_ext e) (fst_ext e) x0)) + (merge e g H Haff)). +apply In_merge_pref_edge. assumption. +intro. destruct (eq_charac _ _ H8); destruct H9. +destruct H5. elim H6. rewrite H5. auto. +elim H7. rewrite H5. auto. +destruct H5. elim H7. rewrite H5. auto. +elim H6. rewrite H5. auto. +assumption. + +intro. unfold Interfere in H8. unfold redirect in H8. +destruct (OTFacts.eq_dec (fst_ext x0) (snd_ext e)); change_rewrite. +destruct H5. elim H7. rewrite H5. auto. +elim H1. rewrite in_pref. destruct H3. exists x1. +assert (eq (x, snd_ext e, Some x1) (snd_ext x0, fst_ext x0, Some x1)) by Eq_eq. +rewrite H9. rewrite <-H3. rewrite edge_comm. change Regs.registers with Regs.t. rewrite <-(edge_eq x0). auto. +destruct (OTFacts.eq_dec (snd_ext x0) (snd_ext e)); change_rewrite. +destruct H5. +elim H1. rewrite in_pref. destruct H3. exists x1. +assert (eq (x, snd_ext e, Some x1) (fst_ext x0, snd_ext x0, Some x1)) by Eq_eq. +rewrite H9. rewrite <-H3. rewrite <-(edge_eq x0). auto. +elim H7. rewrite H5. auto. +generalize (In_merge_edge_inv _ _ _ H Haff H8). intro. +destruct H9. destruct H9. destruct H10. +unfold redirect in H10. destruct (OTFacts.eq_dec (fst_ext x1) (snd_ext e)). +unfold weak_eq in H10; destruct H10; destruct H10; change_rewrite. +destruct H5. elim H6. rewrite H5. auto. +elim H0. rewrite in_pref. destruct H3. exists x2. +assert (eq (x, fst_ext e, Some x2) (snd_ext x0, fst_ext x0, Some x2)) by Eq_eq. +rewrite H13. rewrite edge_comm. rewrite <-H3. change Regs.registers with Regs.t. rewrite <-(edge_eq x0). auto. +destruct H5. +elim H0. rewrite in_pref. destruct H3. exists x2. +assert (eq (x, fst_ext e, Some x2) (fst_ext x0, snd_ext x0, Some x2)) by Eq_eq. +rewrite H13. rewrite <-H3. rewrite <-(edge_eq x0). auto. +elim H6. rewrite H5. auto. +destruct (OTFacts.eq_dec (snd_ext x1) (snd_ext e)). +unfold weak_eq in H10; destruct H10; destruct H10; change_rewrite. +destruct H5. +elim H0. rewrite in_pref. destruct H3. exists x2. +assert (eq (x, fst_ext e, Some x2) (fst_ext x0, snd_ext x0, Some x2)) by Eq_eq. +rewrite H13. rewrite <-H3. rewrite <-(edge_eq x0). auto. +elim H6. rewrite H5. auto. +destruct H5. elim H6. rewrite H5. auto. +elim H0. rewrite in_pref. destruct H3. exists x2. +assert (eq (x, fst_ext e, Some x2) (snd_ext x0, fst_ext x0, Some x2)) by Eq_eq. +rewrite H13. rewrite <-H3. rewrite edge_comm. change Regs.registers with Regs.t. rewrite <-(edge_eq x0). auto. +elim (interf_pref_conflict (fst_ext x0) (snd_ext x0) g). +unfold Prefere, Interfere. split. +destruct H3. exists x2. rewrite <-H3. change Regs.registers with Regs.t. rewrite <-(edge_eq x0). auto. +assert (eq (fst_ext x0, snd_ext x0, None) (fst_ext x1, snd_ext x1, None)). +unfold weak_eq in H10; destruct H10; destruct H10; change_rewrite. +Eq_eq. Eq_comm_eq. rewrite H12. +unfold same_type in H11. destruct H11; destruct H11. +destruct H13. simpl in H13. congruence. +rewrite <-H11. rewrite <-(edge_eq x1). auto. + +unfold redirect in H8. destruct (OTFacts.eq_dec (fst_ext x0) (snd_ext e)); change_rewrite. +destruct H5. elim H7. rewrite H5. auto. +elim H1. rewrite in_pref. destruct H3. exists x1. +assert (eq (x, snd_ext e, Some x1) (fst_ext x0, snd_ext x0, Some x1)) by Eq_comm_eq. +rewrite H9. rewrite <-H3. rewrite <-(edge_eq x0). auto. +destruct (OTFacts.eq_dec (snd_ext x0) (snd_ext e)); change_rewrite. +destruct H5. +elim H1. rewrite in_pref. destruct H3. exists x1. +assert (eq (x, snd_ext e, Some x1) (fst_ext x0, snd_ext x0, Some x1)) by Eq_eq. +rewrite H9. rewrite <-H3. rewrite <-(edge_eq x0). auto. +elim H7. rewrite H5. auto. +unfold Prefere in H8. destruct H8. +apply (move_related_charac2 _ _ (fst_ext x0, snd_ext x0, Some x1)). +unfold aff_edge. exists x1. auto. +assumption. +destruct H5;[left|right]; auto. +Qed. + +Lemma interfere_dec : forall x y g, +Interfere x y g \/ ~Interfere x y g. + +Proof. +intros. unfold Interfere. +destruct (RegRegProps.In_dec (x,y,None) (IE g)). +left. right. assumption. +right. intro. elim n. +rewrite In_graph_interf_edge_in_IE. split. +unfold interf_edge. auto. +assumption. +Qed. + +(* A vertex x different from r which is move-related in g and nonmove-related + in (merge e g p q) is either + 1) interfering with the first endpoint of e and prefering with the second one + 2) interfering with the second endpoint of e and prefering with the first one *) +Lemma move_merge_not_move : forall x e g p q, +~Register.eq x (fst_ext e) -> +~Register.eq x (snd_ext e) -> +move_related (merge e g p q) x = false -> +move_related g x = true -> +(VertexSet.In x (VertexSet.inter (interference_adj (fst_ext e) g) (preference_adj (snd_ext e) g))) \/ +(VertexSet.In x (VertexSet.inter (interference_adj (snd_ext e) g) (preference_adj (fst_ext e) g))). + +Proof. +intros x e g p q H0 H1 H2 H3. generalize I. intro H. +generalize (move_related_charac _ _ H3). intro. +destruct H4. destruct H4. destruct H5. +cut (Interfere (fst_ext (redirect (snd_ext e) (fst_ext e) x0)) + (snd_ext (redirect (snd_ext e) (fst_ext e) x0)) + (merge e g p q)). intro. +unfold Interfere in H7. unfold redirect in H7. +destruct (OTFacts.eq_dec (fst_ext x0) (snd_ext e)); change_rewrite. +destruct H6. +elim H1. rewrite H6. auto. +generalize (In_merge_edge_inv e _ g p q H7). intro. +destruct H8. destruct H8. destruct H9 as [H9 HH9]. +assert (eq (fst_ext e, snd_ext x0, None) (redirect (snd_ext e) (fst_ext e) x1)). +unfold weak_eq; destruct H9; destruct H9; change_rewrite. +apply eq_ordered_eq. split; intuition. change_rewrite. rewrite redirect_weight_eq. +unfold same_type in HH9. destruct HH9. destruct H11. +destruct H12. simpl in H12. congruence. +destruct H11. rewrite <-H11. apply OptionN_as_OT.eq_refl. +Eq_comm_eq. simpl. rewrite redirect_weight_eq. +unfold same_type in HH9. destruct HH9. destruct H11. +destruct H12. simpl in H12. congruence. +destruct H11. rewrite <-H11. apply OptionN_as_OT.eq_refl. +generalize H10. clear H9 HH9 H10. intro H9. +unfold redirect in H9. +destruct (OTFacts.eq_dec (fst_ext x1) (snd_ext e)). +destruct (eq_charac _ _ H9); change_rewrite; destruct H10. +elim (interf_pref_conflict x (snd_ext e) g). +split. +destruct H4. +unfold Prefere. exists x2. +cut (eq (x, snd_ext e, Some x2) x0). intro. +rewrite H12. auto. +rewrite (edge_eq x0). +rewrite edge_comm. apply eq_ordered_eq. constructor; simpl. +split; intuition. +rewrite H4. apply OptionN_as_OT.eq_refl. +unfold Interfere. +cut (eq (x, snd_ext e, None) x1). intro. +rewrite H12. auto. +rewrite (edge_eq x1). +rewrite edge_comm. apply eq_ordered_eq. constructor;simpl. +split; intuition. rewrite H6. auto. +generalize (get_weight_m _ _ (eq_sym H9)). simpl. intro. +rewrite H12. apply OptionN_as_OT.eq_refl. + +elim H0. rewrite H6. auto. +destruct (OTFacts.eq_dec (snd_ext x1) (snd_ext e)). +destruct (eq_charac _ _ H9); change_rewrite. destruct H10. +elim H0. rewrite H6. auto. +destruct H10. + +elim (interf_pref_conflict x (snd_ext e) g). +split. +destruct H4. +unfold Prefere. exists x2. +cut (eq (x, snd_ext e, Some x2) x0). intro. +rewrite H12. auto. +rewrite (edge_eq x0). +rewrite edge_comm. apply eq_ordered_eq. constructor; simpl. +split; intuition. +rewrite H4. apply OptionN_as_OT.eq_refl. +unfold Interfere. +cut (eq (x, snd_ext e, None) x1). intro. +rewrite H12. auto. +rewrite (edge_eq x1). +apply eq_ordered_eq. constructor;simpl. +split; try rewrite H6; intuition. +generalize (get_weight_m _ _ (eq_sym H9)). simpl. intro. +rewrite H12. apply OptionN_as_OT.eq_refl. +left. apply VertexSet.inter_3. +rewrite in_interf. +cut (eq (x, fst_ext e, None) (fst_ext e, snd_ext x0, None)). intro. +rewrite H10. rewrite H9. auto. +rewrite edge_comm. apply eq_ordered_eq. +constructor; simpl. split; intuition. +apply OptionN_as_OT.eq_refl. +destruct H4. +rewrite in_pref. exists x2. +cut (eq (x,snd_ext e, Some x2) x0). intro. +rewrite H10. auto. +rewrite edge_comm. apply eq_ordered_eq. rewrite (edge_eq x0). +constructor; simpl. split; intuition. +auto. auto. rewrite H4. apply OptionN_as_OT.eq_refl. + +destruct (OTFacts.eq_dec (snd_ext x0) (snd_ext e)); change_rewrite. +destruct H6. +generalize (In_merge_edge_inv e _ g p q H7). intro. +destruct H8. destruct H8. destruct H9 as [H9 HH9]. +assert (eq (fst_ext x0, fst_ext e, None) (redirect (snd_ext e) (fst_ext e) x1)). +unfold weak_eq; destruct H9; destruct H9; change_rewrite. +Eq_eq. simpl. rewrite redirect_weight_eq. +unfold same_type in HH9. destruct HH9. destruct H11. +destruct H12. simpl in H12. congruence. +destruct H11. rewrite <-H11. apply OptionN_as_OT.eq_refl. +Eq_comm_eq. simpl. rewrite redirect_weight_eq. +unfold same_type in HH9. destruct HH9. destruct H11. +destruct H12. simpl in H12. congruence. +destruct H11. rewrite <-H11. apply OptionN_as_OT.eq_refl. +generalize H10. clear H9 HH9 H10. intro H9. +unfold redirect in H9. +destruct (OTFacts.eq_dec (fst_ext x1) (snd_ext e)). +destruct (eq_charac _ _ H9); change_rewrite; destruct H10. +elim H0. rewrite H6. auto. +elim (interf_pref_conflict x (snd_ext e) g). +split. +destruct H4. +unfold Prefere. exists x2. +cut (eq (x, snd_ext e, Some x2) x0). intro. +rewrite H12. auto. +rewrite (edge_eq x0). +apply eq_ordered_eq. constructor; simpl. +split; intuition. +auto. +auto. +rewrite H4. apply OptionN_as_OT.eq_refl. +unfold Interfere. +cut (eq (x, snd_ext e, None) x1). intro. +rewrite H12. auto. +rewrite (edge_eq x1). +rewrite edge_comm. apply eq_ordered_eq. constructor;simpl. +split; try rewrite H6; intuition. +generalize (get_weight_m _ _ (eq_sym H9)). simpl. intro. +rewrite H12. apply OptionN_as_OT.eq_refl. + +destruct (OTFacts.eq_dec (snd_ext x1) (snd_ext e)). +destruct (eq_charac _ _ H9); change_rewrite. destruct H10. +elim (interf_pref_conflict x (snd_ext e) g). +split. +destruct H4. +unfold Prefere. exists x2. +cut (eq (x, snd_ext e, Some x2) x0). intro. +rewrite H12. auto. +rewrite (edge_eq x0). +apply eq_ordered_eq. constructor; simpl. +split; intuition. +rewrite H4. apply OptionN_as_OT.eq_refl. +unfold Interfere. +cut (eq (x, snd_ext e, None) x1). intro. +rewrite H12. auto. +rewrite (edge_eq x1). +apply eq_ordered_eq. constructor;simpl. +split; try rewrite H6; intuition. +generalize (get_weight_m _ _ (eq_sym H9)). simpl. intro. +rewrite H12. apply OptionN_as_OT.eq_refl. + +destruct H10. elim H0. rewrite H6. auto. + +left. apply VertexSet.inter_3. +rewrite in_interf. +cut (eq (x, fst_ext e, None) (fst_ext x0, fst_ext e, None)). intro. +rewrite H10. rewrite H9. auto. +apply eq_ordered_eq. +constructor; simpl. split; intuition. +apply OptionN_as_OT.eq_refl. +destruct H4. +rewrite in_pref. exists x2. +cut (eq (x,snd_ext e, Some x2) x0). intro. +rewrite H10. auto. +apply eq_ordered_eq. rewrite (edge_eq x0). +constructor; simpl. split; intuition. +rewrite H4. apply OptionN_as_OT.eq_refl. + +elim H1. rewrite H6. auto. +generalize (In_merge_edge_inv e _ g p q H7). intro. +destruct H8. destruct H8. destruct H9 as [H9 HH9]. +assert (eq (fst_ext x0, snd_ext x0, None) (redirect (snd_ext e) (fst_ext e) x1)). +unfold weak_eq; destruct H9; destruct H9; change_rewrite. +Eq_eq. simpl. rewrite redirect_weight_eq. +unfold same_type in HH9. destruct HH9. destruct H11. +destruct H12. simpl in H12. congruence. +destruct H11. rewrite <-H11. apply OptionN_as_OT.eq_refl. +Eq_comm_eq. simpl. rewrite redirect_weight_eq. +unfold same_type in HH9. destruct HH9. destruct H11. +destruct H12. simpl in H12. congruence. +destruct H11. rewrite <-H11. apply OptionN_as_OT.eq_refl. +generalize H10. clear H9 HH9 H10. intro H9. +unfold redirect in H9. +destruct (OTFacts.eq_dec (fst_ext x1) (snd_ext e)). +destruct (eq_charac _ _ H9); change_rewrite; destruct H10. +destruct H6. elim H0. rewrite H6. auto. +right. apply VertexSet.inter_3. +rewrite in_interf. +cut (eq (x, snd_ext e, None) x1). intro. +rewrite H12. auto. +rewrite edge_comm. apply eq_ordered_eq. +constructor; simpl. split; intuition. +auto. rewrite H6. auto. +generalize (get_weight_m _ _ H9). simpl. intro. +rewrite H12. unfold get_weight. apply OptionN_as_OT.eq_refl. +destruct H4. +rewrite in_pref. exists x2. +cut (eq (x, fst_ext e, Some x2) x0). intro. +rewrite H12. auto. +rewrite edge_comm. apply eq_ordered_eq. rewrite (edge_eq x0). +constructor; simpl. split; intuition. +auto. auto. rewrite H4. apply OptionN_as_OT.eq_refl. + +destruct H6. +right. apply VertexSet.inter_3. +rewrite in_interf. +cut (eq (x, snd_ext e, None) x1). intro. +rewrite H12. auto. +rewrite edge_comm. apply eq_ordered_eq. +constructor; simpl. split; intuition. +auto. rewrite H6. auto. +generalize (get_weight_m _ _ H9). simpl. intro. +rewrite H12. unfold get_weight. apply OptionN_as_OT.eq_refl. +destruct H4. +rewrite in_pref. exists x2. +cut (eq (x, fst_ext e, Some x2) x0). intro. +rewrite H12. auto. +apply eq_ordered_eq. rewrite (edge_eq x0). +constructor; simpl. split; intuition. +auto. auto. rewrite H4. apply OptionN_as_OT.eq_refl. + +elim H0. rewrite H6. auto. +destruct (OTFacts.eq_dec (snd_ext x1) (snd_ext e)). +destruct (eq_charac _ _ H9); change_rewrite; destruct H10. +destruct H6. +right. apply VertexSet.inter_3. +rewrite in_interf. +cut (eq (x, snd_ext e, None) x1). intro. +rewrite H12. auto. +apply eq_ordered_eq. +constructor; simpl. split; intuition. +rewrite H6. auto. auto. +generalize (get_weight_m _ _ H9). simpl. intro. +rewrite H12. unfold get_weight. apply OptionN_as_OT.eq_refl. +destruct H4. +rewrite in_pref. exists x2. +cut (eq (x, fst_ext e, Some x2) x0). intro. +rewrite H12. auto. +apply eq_ordered_eq. rewrite (edge_eq x0). +constructor; simpl. split; intuition. +auto. auto. rewrite H4. apply OptionN_as_OT.eq_refl. + +elim H0. rewrite H6. auto. + +destruct H6. +elim H0. rewrite H6. auto. +right. apply VertexSet.inter_3. +rewrite in_interf. +cut (eq (x, snd_ext e, None) x1). intro. +rewrite H12. auto. +apply eq_ordered_eq. +constructor; simpl. split; intuition. +rewrite H6. auto. auto. +generalize (get_weight_m _ _ H9). simpl. intro. +rewrite H12. unfold get_weight. apply OptionN_as_OT.eq_refl. +destruct H4. +rewrite in_pref. exists x2. +cut (eq (x, fst_ext e, Some x2) x0). intro. +rewrite H12. auto. +rewrite edge_comm. apply eq_ordered_eq. rewrite (edge_eq x0). +constructor; simpl. split; intuition. +auto. auto. rewrite H4. apply OptionN_as_OT.eq_refl. + +elim (interf_pref_conflict (fst_ext x0) (snd_ext x0) g). +split. +unfold Prefere. destruct H4. exists x2. +rewrite <-H4. change Regs.registers with Regs.t. rewrite <-(edge_eq x0). assumption. +unfold Interfere. rewrite H9. assumption. + +(* cut *) +destruct (interfere_dec (fst_ext (redirect (snd_ext e) (fst_ext e) x0)) + (snd_ext (redirect (snd_ext e) (fst_ext e) x0)) + (merge e g p q)). auto. +cut (~eq e x0). intro. +generalize (In_merge_pref_edge e x0 g p q H5 H8 H4 H7). intro. clear H7. +unfold Prefere in H9. destruct H9. unfold redirect in H7. +destruct (OTFacts.eq_dec (fst_ext x0) (snd_ext e)); change_rewrite. +destruct H6. +elim H1. rewrite H6. auto. +assert (move_related (merge e g p q) x = true). +apply move_related_charac2 with (e:= (fst_ext e, snd_ext x0, Some x1)). +exists x1. auto. +auto. +right. change_rewrite. auto. +rewrite H2 in H9. inversion H9. + +destruct (OTFacts.eq_dec (snd_ext x0) (snd_ext e)); change_rewrite. +destruct H6. +assert (move_related (merge e g p q) x = true). +apply move_related_charac2 with (e:= (fst_ext x0, fst_ext e, Some x1)). +exists x1. auto. +auto. +left. change_rewrite. auto. +rewrite H2 in H9. inversion H9. + +elim H1. rewrite H6. auto. + +assert (move_related (merge e g p q) x = true). +apply move_related_charac2 with (e:= (fst_ext x0, snd_ext x0, Some x1)). +exists x1. auto. +auto. +destruct H6;[left|right]; auto. +rewrite H2 in H9. inversion H9. + +intro. destruct (eq_charac _ _ H8); destruct H9. +destruct H6. +elim H0. rewrite H6. auto. +elim H1. rewrite H6. auto. +destruct H6. +elim H1. rewrite H6. auto. +elim H0. rewrite H6. auto. +Qed. + +(* A move-related vertex of g which does not interfere with an endpoint + of e and prefere with the other one is move-related in (merge e g p q) *) +Lemma move_related_merge_true : forall x e g p q, +move_related g x = true -> +~Register.eq x (fst_ext e) -> +~Register.eq x (snd_ext e) -> +~VertexSet.In x (VertexSet.inter (interference_adj (fst_ext e ) g) + (preference_adj (snd_ext e) g)) -> +~VertexSet.In x (VertexSet.inter (interference_adj (snd_ext e) g) + (preference_adj (fst_ext e) g)) -> +move_related (merge e g p q) x = true. + +Proof. +intros x e g p q H0 H1 H2 H3 H4. generalize I. intro H. +case_eq (move_related (merge e g p q) x); intros. +reflexivity. +generalize (move_merge_not_move _ _ _ p q H1 H2 H5 H0). intro. +destruct H6;[elim (H3 H6)|elim (H4 H6)]. +Qed. + +(* A vertex which interferes with an endpoint of e, preferes with + the other one and has a preference degree of one is + nonmove-related in (merge e g p q) *) +Lemma move_related_change_charac : forall x e g p q, +VertexSet.In x (VertexSet.union (VertexSet.inter (interference_adj (fst_ext e) g) + (preference_adj (snd_ext e) g)) + (VertexSet.inter (interference_adj (snd_ext e) g) + (preference_adj (fst_ext e) g))) -> +pref_degree g x = 1 -> +move_related (merge e g p q) x = false. + +Proof. +unfold pref_degree. intros x e g p q H1 H2. generalize I. intro H. +assert (move_related g x = true). +case_eq (move_related g x); auto. intro. +rewrite move_related_card in H0. inversion H0. +unfold pref_degree. congruence. +assert (~Register.eq x (fst_ext e)) as Hfst. +intro. rewrite H3 in H1. +destruct (VertexSet.union_1 H1). +elim (not_in_interf_self (fst_ext e) g). +apply (VertexSet.inter_1 H4). +elim (not_in_pref_self (fst_ext e) g). +apply (VertexSet.inter_2 H4). +assert (~Register.eq x (snd_ext e)) as Hsnd. +intro. rewrite H3 in H1. +destruct (VertexSet.union_1 H1). +elim (not_in_pref_self (snd_ext e) g). +apply (VertexSet.inter_2 H4). +elim (not_in_interf_self (snd_ext e) g). +apply (VertexSet.inter_1 H4). +destruct (VertexSet.union_1 H1). +assert (VertexSet.Equal (preference_adj x g) (VertexSet.singleton (snd_ext e))). +apply cardinal_1_singleton. +apply pref_adj_comm. apply (VertexSet.inter_2 H3). +assumption. + +assert (VertexSet.Equal (preference_adj x (merge e g p q)) VertexSet.empty). +split; intros. +assert (VertexSet.Subset (preference_adj x (merge e g p q)) + (VertexSet.add (fst_ext e) + (VertexSet.remove (snd_ext e) (preference_adj x g)))). +apply preference_adj_merge; assumption. +destruct (Register.eq_dec a (fst_ext e)). +elim (interf_pref_conflict (fst_ext e) x (merge e g p q)). +split. +rewrite in_pref in H5. destruct H5. +unfold Prefere. exists x0. +assert (eq (fst_ext e, x, Some x0) (a, x, Some x0)) by Eq_eq. +rewrite H7. assumption. +unfold Interfere. +assert (eq (fst_ext e, x,None) (redirect (snd_ext e) (fst_ext e) (fst_ext e, x,None))). +apply eq_ordered_eq. +split; simpl; try split; try apply Regs.eq_refl. +unfold redirect; change_rewrite. +destruct (OTFacts.eq_dec (fst_ext e) (snd_ext e)). apply Regs.eq_refl. +destruct (OTFacts.eq_dec x (snd_ext e)). apply Regs.eq_refl. apply Regs.eq_refl. +unfold redirect; change_rewrite; simpl. +destruct (OTFacts.eq_dec (fst_ext e) (snd_ext e)). apply Regs.eq_refl. +destruct (OTFacts.eq_dec x (snd_ext e)). simpl. +elim (not_in_pref_self (snd_ext e) g). +rewrite r in H3. apply (VertexSet.inter_2 H3). apply Regs.eq_refl. +auto. rewrite redirect_weight_eq. simpl. apply OptionN_as_OT.eq_refl. +rewrite H7. apply In_merge_interf_edge. +rewrite <-in_interf. apply interf_adj_comm. apply (VertexSet.inter_1 H3). +unfold interf_edge. auto. + +assert (VertexSet.In a (VertexSet.add (fst_ext e) (VertexSet.remove (snd_ext e) (preference_adj x g)))). +apply H6. assumption. +assert (VertexSet.In a (VertexSet.remove (snd_ext e) (preference_adj x g))). +apply VertexSet.add_3 with (x:=fst_ext e); auto. +rewrite H4 in H8. +assert (~Register.eq a (snd_ext e)). +intro. elim (VertexSet.remove_1 (Register.eq_sym _ _ H9) H8). +generalize (VertexSet.remove_3 H8). intro. +generalize (VertexSet.singleton_1 H10). intro. elim H9. auto. +elim (VertexSet.empty_1 H5). +case_eq (move_related (merge e g p q) x); intros. +elim (move_related_not_empty_pref _ _ H6). assumption. +reflexivity. + +(* symetric part *) +assert (VertexSet.Equal (preference_adj x g) (VertexSet.singleton (fst_ext e))). +apply cardinal_1_singleton. +apply pref_adj_comm. apply (VertexSet.inter_2 H3). +assumption. + +assert (VertexSet.Equal (preference_adj x (merge e g p q)) VertexSet.empty). +split; intros. +assert (VertexSet.Subset (preference_adj x (merge e g p q)) + (VertexSet.add (fst_ext e) + (VertexSet.remove (snd_ext e) (preference_adj x g)))). +apply preference_adj_merge; assumption. +destruct (Register.eq_dec a (fst_ext e)). +elim (interf_pref_conflict (fst_ext e) x (merge e g p q)). +split. +rewrite in_pref in H5. destruct H5. +unfold Prefere. exists x0. +assert (eq (fst_ext e, x, Some x0) (a, x, Some x0)) by Eq_eq. +rewrite H7. assumption. +unfold Interfere. +assert (eq (fst_ext e, x,None) (redirect (snd_ext e) (fst_ext e) (snd_ext e, x,None))). +apply eq_ordered_eq. +split; simpl; try split; auto. +unfold redirect; change_rewrite. +destruct (OTFacts.eq_dec (snd_ext e) (snd_ext e)). apply Regs.eq_refl. +elim n. apply Regs.eq_refl. +unfold redirect; change_rewrite. +destruct (OTFacts.eq_dec (snd_ext e) (snd_ext e)). apply Regs.eq_refl. +elim n. apply Regs.eq_refl. +rewrite redirect_weight_eq. simpl. apply OptionN_as_OT.eq_refl. +rewrite H7. apply In_merge_interf_edge. +rewrite <-in_interf. apply interf_adj_comm. apply (VertexSet.inter_1 H3). +unfold interf_edge. auto. + +assert (VertexSet.In a (VertexSet.add (fst_ext e) (VertexSet.remove (snd_ext e) (preference_adj x g)))). +apply H6. assumption. +assert (VertexSet.In a (VertexSet.remove (snd_ext e) (preference_adj x g))). +apply VertexSet.add_3 with (x:=fst_ext e); auto. +rewrite H4 in H8. +generalize (VertexSet.remove_3 H8). intro. +generalize (VertexSet.singleton_1 H9). intro. elim n. auto. +elim (VertexSet.empty_1 H5). +case_eq (move_related (merge e g p q) x); intros. +elim (move_related_not_empty_pref _ _ H6). assumption. +reflexivity. +Qed. + +(* A move-related vertex of g which is not move-related in + (merge e g p q) has a preference degree of 1 *) +Lemma move_related_dec_1 : forall x e g p q, +~Register.eq x (fst_ext e) -> +~Register.eq x (snd_ext e) -> +move_related g x = true -> +move_related (merge e g p q) x = false -> +pref_degree g x = 1. + +Proof. +unfold pref_degree. intros x e g p q H0 H1 H2 H3. generalize I. intro H. +generalize (preference_adj_merge_2 x e g p q H0 H1). intro. +cut (VertexSet.cardinal (preference_adj x g) <= 1). intro. +cut (VertexSet.cardinal (preference_adj x g) > 0). intro. +intuition. +generalize (move_related_not_empty_pref x g H2). intro. +case_eq (VertexSet.cardinal (preference_adj x g)); intros. +elim H6. apply empty_is_empty_1. apply (cardinal_inv_1 H7). +intuition. +generalize (not_move_related_empty_pref x (merge e g p q) H3). intro. +generalize (move_merge_not_move _ _ _ p q H0 H1 H3 H2). intro. +destruct H6. +assert (VertexSet.cardinal (VertexSet.remove (fst_ext e) (VertexSet.remove (snd_ext e) (preference_adj x g))) <= + (VertexSet.cardinal (preference_adj x (merge e g p q)))). +apply subset_cardinal. assumption. +rewrite H5 in H7. rewrite empty_cardinal in H7. +rewrite remove_cardinal_2 in H7. +rewrite <-remove_cardinal_1 with (x:=snd_ext e). +apply le_n_S. assumption. +apply pref_adj_comm. apply (VertexSet.inter_2 H6). +intro. generalize (VertexSet.remove_3 H8). intro. +elim (interf_pref_conflict (fst_ext e) x g). +split. +rewrite in_pref in H9. destruct H9. +unfold Prefere. exists x0. assumption. +unfold Interfere. +rewrite edge_comm. generalize (VertexSet.inter_1 H6). intro. +rewrite in_interf in H10. assumption. +assert (VertexSet.cardinal (VertexSet.remove (fst_ext e) (VertexSet.remove (snd_ext e) (preference_adj x g))) <= + (VertexSet.cardinal (preference_adj x (merge e g p q)))). +apply subset_cardinal. assumption. +rewrite H5 in H7. rewrite empty_cardinal in H7. +rewrite <-remove_cardinal_2 with (x:=snd_ext e). +rewrite <-remove_cardinal_1 with (x:=fst_ext e). +apply le_n_S. assumption. +apply VertexSet.remove_2. +apply (In_graph_edge_diff_ext e g p). +apply pref_adj_comm. apply (VertexSet.inter_2 H6). +intro. elim (interf_pref_conflict (snd_ext e) x g). +split. +rewrite in_pref in H8. destruct H8. +unfold Prefere. exists x0. assumption. +generalize (VertexSet.inter_1 H6). intro. +unfold Interfere. rewrite edge_comm. rewrite <-in_interf. assumption. +Qed. + +(* Again, meaningful theorem *) +Theorem Merge_move_evolution : forall x e g p q, +~Register.eq x (fst_ext e) -> +~Register.eq x (snd_ext e) -> +((move_related g x = true /\ move_related (merge e g p q) x = false) + <-> + (pref_degree g x = 1 /\ + (VertexSet.In x (VertexSet.inter (interference_adj (fst_ext e) g) (preference_adj (snd_ext e) g)) \/ + VertexSet.In x (VertexSet.inter (interference_adj (snd_ext e) g) (preference_adj (fst_ext e) g))))). + +Proof. +split; intros. +destruct H1. +split. apply (move_related_dec_1 x e g p q); auto. + apply (move_merge_not_move x e g p q); auto. +destruct H1. +split. case_eq (move_related g x); auto. intro. rewrite move_related_card in H3; congruence. + apply move_related_change_charac; auto. + destruct H2;[apply VertexSet.union_2|apply VertexSet.union_3]; auto. +Qed. diff --git a/backend/Merge_WL.v b/backend/Merge_WL.v new file mode 100755 index 00000000..d9b27e90 --- /dev/null +++ b/backend/Merge_WL.v @@ -0,0 +1,818 @@ +Require Import FSets. +Require Import InterfGraphMapImp. +Require Import Simplify_WL. +Require Import IRC_graph. +Require Import WS. +Require Import Edges. +Require Import Affinity_relation. +Require Import Interference_adjacency. +Require Import Remove_Vertex_WL. +Require Import Merge_Degree. + +Import Edge RegFacts Props OTFacts. + +Definition classify x g' palette wl := +if is_precolored x g' then wl else +let a := get_spillWL wl in +let b := get_freezeWL wl in +let c := get_simplifyWL wl in +let d := get_movesWL wl in +if has_low_degree g' palette x then + if move_related g' x then (a, VertexSet.add x b, c, d) + else (a, VertexSet.remove x b, VertexSet.add x c, d) +else (VertexSet.add x a, VertexSet.remove x b, c, d). + +Definition merge_wl3 e ircg g' (p : In_graph_edge e (irc_g ircg)) (q : aff_edge e) := +let wl := (irc_wl ircg) in +let g := (irc_g ircg) in +let palette := (pal ircg) in +let k := (irc_k ircg) in +let simplifyWL := get_simplifyWL wl in +let freezeWL := get_freezeWL wl in +let spillWL := get_spillWL wl in +let movesWL := get_movesWL wl in +let pre := precolored g in +let iadj_fst_ := interference_adj (fst_ext e) g in +let iadj_fst := VertexSet.diff iadj_fst_ pre in +let iadj_snd_ := interference_adj (snd_ext e) g in +let iadj_snd := VertexSet.diff iadj_snd_ pre in +let padj_fst_ := preference_adj (fst_ext e) g in +let padj_fst := VertexSet.diff padj_fst_ pre in +let padj_snd_ := preference_adj (snd_ext e) g in +let padj_snd := VertexSet.diff padj_snd_ pre in +let iadj_interf := VertexSet.inter iadj_fst iadj_snd in +let iadj_padj_interf := VertexSet.union + (VertexSet.inter padj_fst iadj_snd) + (VertexSet.inter padj_snd iadj_fst) in +let N := VertexSet.filter + (fun x => eq_K k (VertexSet.cardinal (interference_adj x g))) + iadj_interf in +let (mr, nmr) := VertexSet.partition (move_related g) N in +let M := VertexSet.filter + (fun x => eq_K 1 (VertexSet.cardinal (preference_adj x g)) && has_low_degree g k x) + iadj_padj_interf in +let simplifyWL_ := VertexSet.union simplifyWL nmr in +let simplifyWL' := VertexSet.union simplifyWL_ M in +let freezeWL__ := VertexSet.diff freezeWL M in +let freezeWL_ := VertexSet.union freezeWL__ mr in +let freezeWL' := VertexSet.remove (fst_ext e) (VertexSet.remove (snd_ext e) freezeWL_) in +let spillWL_ := VertexSet.diff spillWL N in +let spillWL' := VertexSet.remove (fst_ext e) (VertexSet.remove (snd_ext e) spillWL_) in +let movesWL' := AE_merge_up (preference_adj (fst_ext e) g') + padj_fst_ padj_snd_ e movesWL in +classify (fst_ext e) g' k (spillWL', freezeWL', simplifyWL', movesWL'). + + +Lemma ws_merge3 : forall e ircg H H0, +WS_properties (merge e (irc_g ircg) H H0) + (VertexSet.cardinal (pal ircg)) + (merge_wl3 e ircg (merge e (irc_g ircg) H H0) H H0). + +Proof. +intros e ircg HH HH0. +unfold merge_wl3. unfold WS_properties. rewrite <-(Hk ircg) in *. +set (wl := irc_wl ircg) in *. +set (g := irc_g ircg) in *. +set (g' := merge e g HH HH0) in *. +set (palette := (pal ircg)) in *. +set (simplify := get_simplifyWL wl) in *. +set (freeze := get_freezeWL wl) in *. +set (spill := get_spillWL wl) in *. +set (moves := get_movesWL wl) in *. +set (k := VertexSet.cardinal palette) in *. +set (pre := precolored g) in *. +set (iadj_fst_ := interference_adj (fst_ext e) g) in *. +set (iadj_snd_ := interference_adj (snd_ext e) g) in *. +set (padj_fst_ := preference_adj (fst_ext e) g) in *. +set (padj_snd_ := preference_adj (snd_ext e) g) in *. +set (iadj_fst := VertexSet.diff iadj_fst_ pre) in *. +set (iadj_snd := VertexSet.diff iadj_snd_ pre) in *. +set (padj_fst := VertexSet.diff padj_fst_ pre) in *. +set (padj_snd := VertexSet.diff padj_snd_ pre) in *. +set (inter_interf := VertexSet.inter iadj_fst iadj_snd) in *. +set (iadj_padj_interf := VertexSet.union + (VertexSet.inter padj_fst iadj_snd) + (VertexSet.inter padj_snd iadj_fst)) in *. +set (N := VertexSet.filter + (fun x => eq_K k (VertexSet.cardinal (interference_adj x g))) + inter_interf) in *. +set (M := VertexSet.filter + (fun x => eq_K 1 (VertexSet.cardinal (preference_adj x g)) + && has_low_degree g k x) + iadj_padj_interf) in *. +set (mrnmr := VertexSet.partition (move_related g) N) in *. +case_eq mrnmr. intros mr nmr Hmr. unfold mrnmr in Hmr. +set (simplifyWL_ := VertexSet.union simplify nmr) in *. +set (simplifyWL' := VertexSet.union simplifyWL_ M) in *. +set (freezeWL__ := VertexSet.diff freeze M) in *. +set (freezeWL_ := VertexSet.union freezeWL__ mr) in *. +set (freezeWL' := VertexSet.remove (fst_ext e) (VertexSet.remove (snd_ext e) freezeWL_)) in *. +set (spillWL_ := VertexSet.diff spill N) in *. +set (spillWL' := VertexSet.remove (fst_ext e) (VertexSet.remove (snd_ext e) spillWL_)) in *. +set (movesWL' := AE_merge_up (preference_adj (fst_ext e) g') padj_fst_ padj_snd_ e moves) in *. + +unfold get_spillWL. unfold get_freezeWL. unfold get_simplifyWL. unfold get_movesWL. simpl. +generalize (HWS_irc ircg). intro HWS. rewrite <-(Hk ircg) in HWS. +generalize HH0. intro Haffa. destruct Haffa as [x2 Haffa]. + +cut (VertexSet.Equal mr (VertexSet.filter (move_related g) N)). +intro HMR. +cut (VertexSet.Equal nmr (VertexSet.filter (fun x => negb (move_related g x)) N)). +intro HNMR. + +split. +intro x. +(* spillWL for x = fst_ext e *) +destruct (Register.eq_dec x (fst_ext e)); split; intros. +(* spillWL => for x = fst_ext e *) +unfold classify in H. +unfold get_simplifyWL, get_freezeWL, get_spillWL, get_movesWL in H. simpl in H. +(* fst_ext e is precolored => False *) +case_eq (is_precolored (fst_ext e) g'); intros; rewrite H0 in H. +simpl in H. +unfold spillWL' in H. generalize (VertexSet.remove_3 (VertexSet.remove_3 H)). intro. +unfold spillWL_ in H1. generalize (VertexSet.diff_1 H1). intro. +generalize (In_spill_props _ _ _ _ _ _ _ _ H2 (refl_equal _) HWS). intro. +elim (proj2 (proj2 H3)). +apply VertexSet.remove_3 with (x:=snd_ext e). +rewrite <-(precolored_merge e (irc_g ircg) HH HH0). rewrite precolored_equiv. split. +rewrite is_precolored_ext with (y := fst_ext e); eassumption. +rewrite In_merge_vertex. split. intuition. +intro. elim (In_graph_edge_diff_ext e g). assumption. +rewrite <-e0. rewrite <-H4. auto. +(* fst_ext e is not precolored *) +case_eq (has_low_degree g' k (fst_ext e)); intros; rewrite H1 in H. +(* fst_ext e is of low-degree in g' *) +case_eq (move_related g' (fst_ext e)); intros; rewrite H2 in H; +simpl in H; unfold spillWL' in H; elim (VertexSet.remove_1 (Register.eq_sym _ _ e0) H). +(* fst_ext e is of high-degree in g' *) +simpl in H. split. +rewrite (compat_bool_low _ _ _ _ e0). assumption. +split. unfold g'. rewrite In_merge_vertex. split. +rewrite e0. apply (proj1 (In_graph_edge_in_ext _ _ HH)). +intro. elim (In_graph_edge_diff_ext e g). assumption. +rewrite <-e0. rewrite <-H2. auto. +rewrite precolored_equiv. intro. destruct H2. +rewrite (is_precolored_ext _ _ g' e0) in H2. rewrite H0 in H2. inversion H2. +(* spillWL <= for x = fst_ext e *) +destruct H. destruct H0. +unfold classify, get_spillWL, get_simplifyWL, get_freezeWL, get_movesWL. simpl. +case_eq (is_precolored (fst_ext e) g'); intros; simpl. +(* fst_ext e is precolored => False *) +elim H1. rewrite precolored_equiv. split. +rewrite (is_precolored_ext _ _ g' e0). assumption. assumption. +(* fst_ext e is not precolored *) +case_eq (has_low_degree g' k (fst_ext e)); intros; simpl. +(* fst_ext e is of low-degree in g' *) +rewrite (compat_bool_low _ _ _ _ e0) in H. rewrite H in H3. inversion H3. +(* fst_ext e is of high-degree in g' *) +apply VertexSet.add_1. intuition. + +(* spillWL : x is different from fst_ext e *) +assert (VertexSet.In x spillWL'). +unfold classify, get_simplifyWL, get_movesWL, get_freezeWL, get_spillWL in H. simpl in H. +case_eq (is_precolored (fst_ext e) g'); simpl in H; intro; rewrite H0 in H. +assumption. +case_eq (has_low_degree g' k (fst_ext e)); intros; rewrite H1 in H; simpl in H. +case_eq (move_related g' (fst_ext e)); intros; rewrite H2 in H; simpl in H; assumption. +apply VertexSet.add_3 with (x:=fst_ext e). auto. assumption. +generalize H0. clear H H0. intro H. +unfold spillWL' in H. generalize (VertexSet.remove_3 H). clear H. intro H. +assert (~Register.eq x (snd_ext e)). +intro. elim (VertexSet.remove_1 (Register.eq_sym _ _ H0) H). +generalize (VertexSet.remove_3 H). clear H. intro H. +unfold spillWL_ in H. +generalize (VertexSet.diff_1 H). intro. +generalize (VertexSet.diff_2 H). clear H. intro H. +generalize (In_spill_props _ _ _ _ _ _ _ _ H1 (refl_equal _) HWS). intro. +destruct H2. destruct H3. fold palette in H2. fold k in H2. +split. +case_eq (has_low_degree g' k x); intros. +generalize (merge_dec_interf _ _ _ _ HH HH0 H2 H5 n H0). intro. +generalize (merge_dec_K _ _ _ _ HH HH0 n H0 H2 H5). intro. +destruct H6. elim H. unfold N. +apply VertexSet.filter_3. +apply eq_K_compat. +unfold inter_interf. apply VertexSet.inter_3. +apply VertexSet.diff_3; assumption. +apply VertexSet.diff_3; assumption. +apply eq_K_1. assumption. +reflexivity. +split. unfold g'. rewrite In_merge_vertex. split; assumption. +unfold g'. rewrite precolored_merge. +intro. elim H4. apply (VertexSet.remove_3 H5). +(* spillWL <= for x different from fst_ext e *) +destruct H. destruct H0. +assert (VertexSet.In x spillWL'). +unfold spillWL'. apply VertexSet.remove_2. auto. +apply VertexSet.remove_2. intro. +unfold g' in H0. rewrite In_merge_vertex in H0. destruct H0. elim H3. auto. +unfold spillWL_. apply VertexSet.diff_3. +WS_apply HWS. +split. +apply (merge_low_1 g e x k HH0 HH). assumption. +intro. rewrite H2 in H0. +unfold g' in H0. rewrite In_merge_vertex in H0. destruct H0. elim H3. auto. +assumption. +split. unfold g' in H0. rewrite In_merge_vertex in H0. intuition. +unfold g' in H1. rewrite precolored_merge in H1. +intro. elim H1. apply VertexSet.remove_2. intro. +rewrite <-H3 in H0. +unfold g' in H0. rewrite In_merge_vertex in H0. destruct H0. elim H4. auto. +assumption. +intro. unfold N in H2. +generalize (VertexSet.filter_1 (eq_K_compat k g) H2). intro. +generalize (VertexSet.filter_2 (eq_K_compat k g) H2). clear H2. intro. +generalize (eq_K_2 _ _ H2). clear H2. intro H2. +assert (has_low_degree g' k x = true). +apply merge_dec_low. auto. +generalize (VertexSet.inter_1 H3). intro. apply (VertexSet.diff_1 H4). +generalize (VertexSet.inter_2 H3). intro. apply (VertexSet.diff_1 H4). +rewrite H in H4. inversion H4. +unfold classify, get_simplifyWL, get_freezeWL, get_movesWL, get_spillWL. simpl. +case_eq (is_precolored (fst_ext e) g'); intros. +assumption. +case_eq (has_low_degree g' k (fst_ext e)); intros. +case_eq (move_related g' (fst_ext e)); intros. +assumption. +assumption. +simpl. apply VertexSet.add_2. assumption. + +(* freezeWL *) +split. +intro x. destruct (Register.eq_dec x (fst_ext e)). +(* x = fst_ext e *) +split; intro. +(* => *) +unfold classify, get_spillWL, get_simplifyWL, get_movesWL, get_freezeWL in H. simpl in H. +case_eq (is_precolored (fst_ext e) g'); intros; rewrite H0 in H; simpl in H. +unfold freezeWL' in H. elim (VertexSet.remove_1 (Register.eq_sym _ _ e0) H). +case_eq (has_low_degree g' k (fst_ext e)); intros; rewrite H1 in H; simpl in H. +split. +rewrite (compat_bool_low _ _ _ _ e0). assumption. +case_eq (move_related g' (fst_ext e)); intros; rewrite H2 in H; simpl in H. +split. +rewrite (compat_bool_move _ _ _ e0). assumption. +rewrite precolored_equiv. intro. destruct H3. +rewrite (is_precolored_ext _ _ g' e0) in H3. rewrite H0 in H3. inversion H3. +elim (VertexSet.remove_1 (Register.eq_sym _ _ e0) H). +elim (VertexSet.remove_1 (Register.eq_sym _ _ e0) H). +(* <= *) +destruct H. destruct H0. +unfold classify, get_simplifyWL, get_spillWL, get_freezeWL, get_movesWL. simpl. +case_eq (is_precolored (fst_ext e) g'); intros; simpl. +elim H1. rewrite precolored_equiv. split. +rewrite (is_precolored_ext _ _ g' e0). assumption. +apply move_related_in_graph. assumption. +case_eq (has_low_degree g' k (fst_ext e)); intros; simpl. +case_eq (move_related g' (fst_ext e)); intros; simpl. +apply VertexSet.add_1. intuition. +rewrite (compat_bool_move _ _ _ e0) in H0. rewrite H0 in H4. inversion H4. +rewrite (compat_bool_low _ _ _ _ e0) in H. rewrite H3 in H. inversion H. + +(* x <> fst_ext e => *) +split; intro. +assert (VertexSet.In x freezeWL'). +unfold classify, get_spillWL, get_simplifyWL, get_freezeWL, get_movesWL in H. simpl in H. +case_eq (is_precolored (fst_ext e) g'); intros; rewrite H0 in H. +assumption. +case_eq (has_low_degree g' k (fst_ext e)); intros; rewrite H1 in H. +case_eq (move_related g' (fst_ext e)); intros; rewrite H2 in H. +simpl in H. apply VertexSet.add_3 with (x:= fst_ext e). auto. assumption. +simpl in H. apply (VertexSet.remove_3 H). +simpl in H. apply (VertexSet.remove_3 H). +generalize H0. clear H H0. intro H. +unfold freezeWL' in H. +generalize (VertexSet.remove_3 H). clear H. intro. +assert (~Register.eq x (snd_ext e)). +intro. elim (VertexSet.remove_1 (Register.eq_sym _ _ H0) H). +generalize (VertexSet.remove_3 H). clear H. intro. +unfold freezeWL_ in H. +destruct (VertexSet.union_1 H). +unfold freezeWL__ in H1. +generalize (VertexSet.diff_1 H1). intro. +generalize (VertexSet.diff_2 H1). clear H1. intro. +unfold freeze in H2. +generalize (In_freeze_props _ _ _ _ _ _ _ _ H2 (refl_equal _) HWS). intro. +destruct H3. destruct H4. destruct H5. +fold palette in H3. fold k in H3. +split. +apply low_dec; assumption. +split. +case_eq (move_related g' x); intros. +reflexivity. +elim H1. unfold M. +apply VertexSet.filter_3. +apply compat_move_up. +Require Import Merge_Move. +Require Import Graph_Facts. +generalize (move_merge_not_move x e g HH HH0 n H0 H7 H4). intro. +destruct H8. +apply VertexSet.union_3. +apply VertexSet.inter_3. +apply VertexSet.diff_3. apply (VertexSet.inter_2 H8). assumption. +apply VertexSet.diff_3. apply (VertexSet.inter_1 H8). assumption. +apply VertexSet.union_2. +apply VertexSet.inter_3. +apply VertexSet.diff_3. apply (VertexSet.inter_2 H8). assumption. +apply VertexSet.diff_3. apply (VertexSet.inter_1 H8). assumption. +generalize (move_related_dec_1 x e g HH HH0 n H0 H4 H7). intro. +unfold pref_degree in H8. rewrite (eq_K_1 _ _ H8). simpl. assumption. +unfold g'. rewrite precolored_merge. intro. elim H6. +apply (VertexSet.remove_3 H7). + +(* x in mr *) +rewrite HMR in H1. +generalize (VertexSet.filter_1 (compat_bool_move _) H1). intro. +generalize (VertexSet.filter_2 (compat_bool_move _) H1). clear H1. intro H1. +unfold N in H2. +generalize (VertexSet.filter_1 (eq_K_compat _ _) H2). intro. +generalize (VertexSet.filter_2 (eq_K_compat _ _) H2). clear H2. intro. +generalize (eq_K_2 _ _ H2). clear H2. intro. +split. +apply merge_dec_low. auto. +apply (VertexSet.diff_1 (VertexSet.inter_1 H3)). +apply (VertexSet.diff_1 (VertexSet.inter_2 H3)). +split. +apply move_related_merge_true; auto. +intro. +generalize (VertexSet.inter_1 H4). intro. +generalize (VertexSet.inter_2 H4). clear H4. intro. +generalize (VertexSet.inter_1 H3). intro. +generalize (VertexSet.inter_2 H3). clear H3. intro. +generalize (VertexSet.diff_1 H3). clear H3. intro. +generalize (VertexSet.diff_1 H6). intro. +generalize (VertexSet.diff_2 H6). clear H6. intro. +elim (interf_pref_conflict x (snd_ext e) g). +split. +rewrite in_pref in H4. destruct H4. +unfold Prefere. exists x0. assumption. +unfold Interfere. rewrite <-in_interf. assumption. +intro. +generalize (VertexSet.inter_1 H4). intro. +generalize (VertexSet.inter_2 H4). clear H4. intro. +generalize (VertexSet.inter_1 H3). intro. +generalize (VertexSet.inter_2 H3). clear H3. intro. +generalize (VertexSet.diff_1 H3). clear H3. intro. +generalize (VertexSet.diff_1 H6). intro. +generalize (VertexSet.diff_2 H6). clear H6. intro. +elim (interf_pref_conflict x (fst_ext e) g). +split. +rewrite in_pref in H4. destruct H4. +unfold Prefere. exists x0. assumption. +unfold Interfere. unfold iadj_fst_ in H7. rewrite in_interf in H7. auto. +unfold g'. rewrite precolored_merge. +intro. elim (VertexSet.diff_2 (VertexSet.inter_1 H3)). +apply (VertexSet.remove_3 H4). + +(* <= *) +assert (VertexSet.In x freezeWL'). +destruct H. destruct H0. +unfold freezeWL'. +apply VertexSet.remove_2. auto. +apply VertexSet.remove_2. +intro. elim (not_in_merge_snd e g HH HH0). +apply move_related_in_graph. rewrite (compat_bool_move _ _ _ H2). assumption. +unfold freezeWL_. +case_eq (has_low_degree g k x); intros. +apply VertexSet.union_2. +unfold freezeWL__. +apply VertexSet.diff_3. +WS_apply HWS. +split. assumption. +split. +apply (move_related_merge_move_related g e x HH HH0). assumption. +unfold g' in H1. rewrite precolored_merge in H1. intro. elim H1. +apply VertexSet.remove_2. intro. elim (not_in_merge_snd _ _ HH HH0). +apply move_related_in_graph. rewrite (compat_bool_move _ _ _ H4). assumption. assumption. +intro. unfold M in H3. +generalize (VertexSet.filter_2 (compat_move_up _ _) H3). intro. +assert (move_related g' x = false). +apply move_related_change_charac. +generalize (VertexSet.filter_1 (compat_move_up _ _) H3). intro. +unfold iadj_padj_interf in H5. +destruct (VertexSet.union_1 H5). +apply VertexSet.union_3. +apply VertexSet.inter_3. +apply (VertexSet.diff_1 (VertexSet.inter_2 H6)). +apply (VertexSet.diff_1 (VertexSet.inter_1 H6)). +apply VertexSet.union_2. +apply VertexSet.inter_3. +apply (VertexSet.diff_1 (VertexSet.inter_2 H6)). +apply (VertexSet.diff_1 (VertexSet.inter_1 H6)). +case_eq (eq_K 1 (VertexSet.cardinal (preference_adj x g))); intros. +symmetry. apply (eq_K_2 _ _ H5). rewrite H5 in H4. inversion H4. +rewrite H0 in H5. inversion H5. +apply VertexSet.union_3. +rewrite HMR. apply VertexSet.filter_3. +apply compat_bool_move. +unfold N. apply VertexSet.filter_3. +apply eq_K_compat. +generalize (merge_dec_interf x e k g HH HH0 H2 H n). intro. +destruct H3. +intro. elim (not_in_merge_snd e g HH HH0). apply move_related_in_graph. +rewrite (compat_bool_move _ _ _ (Register.eq_sym _ _ H3)). assumption. +apply VertexSet.inter_3. +apply VertexSet.diff_3. assumption. unfold pre. +unfold g' in H1. rewrite precolored_merge in H1. intro. elim H1. +apply VertexSet.remove_2. intro. +elim (not_in_merge_snd _ _ HH HH0). rewrite H6. apply move_related_in_graph. assumption. +assumption. +apply VertexSet.diff_3. assumption. unfold pre. +unfold g' in H1. rewrite precolored_merge in H1. intro. elim H1. +apply VertexSet.remove_2. intro. +elim (not_in_merge_snd _ _ HH HH0). rewrite H6. apply move_related_in_graph. assumption. +assumption. +apply eq_K_1. apply (merge_dec_K x e k g HH HH0); auto. +intro. elim (not_in_merge_snd e g HH HH0). +apply move_related_in_graph. rewrite (compat_bool_move _ _ _ H3) in H0. assumption. +apply (move_related_merge_move_related g e x HH HH0). assumption. + +unfold classify, get_spillWL, get_freezeWL, get_simplifyWL, get_movesWL; simpl. +case_eq (is_precolored (fst_ext e) g'); intro; simpl. +assumption. +case_eq (has_low_degree g' k (fst_ext e)); intro; simpl. +case_eq (move_related g' (fst_ext e)); intro; simpl. +apply VertexSet.add_2. assumption. +apply VertexSet.remove_2. auto. assumption. +apply VertexSet.remove_2. auto. assumption. + +(* simplifyWL *) +split. intro x. +(* x = fst_ext e *) +destruct (Register.eq_dec x (fst_ext e)). +(* => *) +split; intros. +unfold classify, get_spillWL, get_simplifyWL, get_movesWL, get_freezeWL in H. simpl in H. +case_eq (is_precolored (fst_ext e) g'); intros; rewrite H0 in H. +simpl in H. unfold simplifyWL' in H. +destruct (VertexSet.union_1 H). +unfold simplifyWL_ in H1. +destruct (VertexSet.union_1 H1). +generalize (In_simplify_props _ _ _ _ _ _ _ _ H2 (refl_equal _) HWS). intro. +destruct H3. destruct H4. destruct H5. +elim H6. apply VertexSet.remove_3 with (x:=snd_ext e). +rewrite <-(precolored_merge e (irc_g ircg) HH HH0). rewrite precolored_equiv. +rewrite (is_precolored_ext _ _ (merge e (irc_g ircg) HH HH0) e0). split. eassumption. +rewrite In_merge_vertex. split. assumption. +intro. elim (In_graph_edge_diff_ext e g). assumption. +rewrite <-H7. assumption. +rewrite HNMR in H2. +generalize (VertexSet.filter_1 (compat_not_compat (compat_bool_move _)) H2). intro. +unfold N in H3. +generalize (VertexSet.filter_1 (eq_K_compat _ _) H3). intro. +generalize (VertexSet.inter_1 H4). intro. +generalize (VertexSet.inter_2 H4). clear H4. intro. +elim (VertexSet.diff_2 H5). unfold pre. +apply VertexSet.remove_3 with (x:=snd_ext e). +rewrite <-(precolored_merge e g HH HH0). rewrite precolored_equiv. +split. rewrite (is_precolored_ext _ _ (merge e g HH HH0) e0). eassumption. +rewrite In_merge_vertex. split. +rewrite e0. apply (proj1 (In_graph_edge_in_ext _ _ HH)). +intro. elim (not_in_interf_self (snd_ext e) g). rewrite H6 in H4. +apply (VertexSet.diff_1 H4). +unfold M in H1. +generalize (VertexSet.filter_1 (compat_move_up _ _) H1). intro. +unfold iadj_padj_interf in H2. +destruct (VertexSet.union_1 H2). +generalize (VertexSet.inter_1 H3). intro. generalize (VertexSet.diff_2 H4). intro. +elim H5. unfold pre. +apply VertexSet.remove_3 with (x:=snd_ext e). rewrite <-(precolored_merge e g HH HH0). +rewrite precolored_equiv. split. +rewrite (is_precolored_ext _ _ (merge e g HH HH0) e0). eassumption. +rewrite e0. rewrite In_merge_vertex. split. +apply (proj1 (In_graph_edge_in_ext _ _ HH)). +intro. elim (In_graph_edge_diff_ext e g HH). auto. +generalize (VertexSet.inter_1 H3). intro. generalize (VertexSet.diff_2 H4). intro. +elim H5. unfold pre. +apply VertexSet.remove_3 with (x:=snd_ext e). rewrite <-(precolored_merge e g HH HH0). +rewrite precolored_equiv. split. +rewrite (is_precolored_ext _ _ (merge e g HH HH0) e0). eassumption. +rewrite e0. rewrite In_merge_vertex. split. +apply (proj1 (In_graph_edge_in_ext _ _ HH)). +intro. elim (In_graph_edge_diff_ext e g HH). auto. + +(* fst_ext e is not precolored *) +case_eq (has_low_degree g' k (fst_ext e)); intros; rewrite H1 in H. +case_eq (move_related g' (fst_ext e)); intros; rewrite H2 in H. +simpl in H. unfold simplifyWL' in H. +destruct (VertexSet.union_1 H). +unfold simplifyWL_ in H3. +destruct (VertexSet.union_1 H3). +assert (move_related g x = false). +apply (proj1 (proj2 (In_simplify_props _ _ _ _ _ _ _ _ H4 (refl_equal _) HWS))). +assert (move_related g x = true). +rewrite (compat_bool_move _ _ _ e0). +apply (proj1 (Aff_edge_aff _ _ HH HH0)). +rewrite H5 in H6. inversion H6. +rewrite HNMR in H4. +generalize (VertexSet.filter_1 (compat_not_compat (compat_bool_move _)) H4). intro. +generalize (VertexSet.filter_2 (compat_not_compat (compat_bool_move _)) H4). clear H4. intro H4. +unfold N in H5. +generalize (VertexSet.filter_1 (eq_K_compat _ _) H5). intro. +generalize (VertexSet.filter_2 (eq_K_compat _ _) H5). clear H5. intro. +generalize (eq_K_2 _ _ H5). clear H5. intro. +assert (move_related g x = true). +rewrite (compat_bool_move _ _ _ e0). +apply (proj1 (Aff_edge_aff _ _ HH HH0)). +rewrite H7 in H4. inversion H4. +unfold M in H3. +generalize (VertexSet.filter_1 (compat_move_up _ _) H3). intro. +unfold iadj_padj_interf in H4. +destruct (VertexSet.union_1 H4). +elim (not_in_pref_self (fst_ext e) g). rewrite e0 in H5. +apply (VertexSet.diff_1 (VertexSet.inter_1 H5)). +elim (not_in_interf_self (fst_ext e) g). rewrite e0 in H5. +apply (VertexSet.diff_1 (VertexSet.inter_2 H5)). +simpl in H. +split. rewrite (compat_bool_low _ _ _ _ e0). assumption. +split. rewrite (compat_bool_move _ _ _ e0). assumption. +split. unfold g'. rewrite In_merge_vertex. split. +rewrite e0. apply (proj1 (In_graph_edge_in_ext e g HH)). +intro. elim (In_graph_edge_diff_ext e g HH). rewrite <-H3. assumption. + +rewrite precolored_equiv. intro. destruct H3. +rewrite (is_precolored_ext _ _ g' e0) in H3. rewrite H0 in H3. inversion H3. +simpl in H. +unfold simplifyWL' in H. +destruct (VertexSet.union_1 H). +unfold simplifyWL_ in H2. +destruct (VertexSet.union_1 H2). +generalize (In_simplify_props _ _ _ _ _ _ _ _ H3 (refl_equal _) HWS). intro. +destruct H4. destruct H5. destruct H6. +generalize (proj1 (Aff_edge_aff _ _ HH HH0)). intro. +rewrite (compat_bool_move _ _ _ e0) in H5. fold g in H5. rewrite H5 in H8. inversion H8. +rewrite HNMR in H3. +generalize (VertexSet.filter_2 (compat_not_compat (compat_bool_move _)) H3). intro. +generalize (proj1 (Aff_edge_aff _ _ HH HH0)). intro. +rewrite (compat_bool_move _ _ _ e0) in H4. rewrite H5 in H4. inversion H4. +unfold M in H2. +generalize (VertexSet.filter_1 (compat_move_up _ _) H2). intro. +unfold iadj_padj_interf in H3. +destruct (VertexSet.union_1 H3). +elim (not_in_pref_self (fst_ext e) g). rewrite e0 in H4. +apply (VertexSet.diff_1 (VertexSet.inter_1 H4)). +elim (not_in_interf_self (fst_ext e) g). rewrite e0 in H4. +apply (VertexSet.diff_1 (VertexSet.inter_2 H4)). + +(* <= *) +destruct H. destruct H0. destruct H1. +unfold classify, get_spillWL, get_simplifyWL, get_freezeWL, get_movesWL. simpl. +case_eq (is_precolored (fst_ext e) g'); intros; simpl. +rewrite precolored_equiv in H2. elim H2. +split. rewrite (is_precolored_ext _ _ g' e0). assumption. assumption. +case_eq (has_low_degree g' k (fst_ext e)); intros; simpl. +case_eq (move_related g' (fst_ext e)); intros; simpl. +rewrite (compat_bool_move _ _ _ e0) in H0. rewrite H0 in H5. inversion H5. +apply VertexSet.add_1. intuition. +rewrite (compat_bool_low _ _ _ _ e0) in H. rewrite H in H4. inversion H4. + +(* x <> fst_ext e *) +(* simplifyWL => *) +split; intros. +assert (VertexSet.In x simplifyWL'). +unfold classify, get_spillWL, get_simplifyWL, get_movesWL, get_freezeWL in H. simpl in H. +case_eq (is_precolored (fst_ext e) g'); intros; rewrite H0 in H. +assumption. +case_eq (has_low_degree g' k (fst_ext e)); intros; rewrite H1 in H. +case_eq (move_related g' (fst_ext e)); intros; rewrite H2 in H. +assumption. +simpl in H. apply VertexSet.add_3 with (x:=fst_ext e). auto. assumption. +assumption. +generalize H0. clear H H0. intro H. +unfold simplifyWL' in H. +destruct (VertexSet.union_1 H). +unfold simplifyWL_ in H0. +destruct (VertexSet.union_1 H0). +generalize (In_simplify_props _ _ _ _ _ _ _ _ H1 (refl_equal _) HWS). intro. +destruct H2. destruct H3. destruct H4. +split. +apply low_dec; auto. +intro. generalize (proj2 (Aff_edge_aff e g HH HH0)). intro. +rewrite (compat_bool_move _ _ _ H6) in H3. fold g in H3. +rewrite H3 in H7. inversion H7. +split. +apply move_merge_false; auto. +split. unfold g'. rewrite In_merge_vertex. split. assumption. +intro. generalize (proj2 (Aff_edge_aff e g HH HH0)). intro. +rewrite (compat_bool_move _ _ _ H6) in H3. fold g in H3. +rewrite H3 in H7. inversion H7. +unfold g'. rewrite precolored_merge. intro. elim H5. +apply (VertexSet.remove_3 H6). +rewrite HNMR in H1. +generalize (VertexSet.filter_1 (compat_not_compat (compat_bool_move _)) H1). intro. +generalize (VertexSet.filter_2 (compat_not_compat (compat_bool_move _)) H1). clear H1. intro. +unfold N in H2. +generalize (VertexSet.filter_1 (eq_K_compat _ _) H2). intro. +generalize (VertexSet.filter_2 (eq_K_compat _ _) H2). clear H2. intro. +generalize (eq_K_2 _ _ H2). clear H2. intro. +split. +apply merge_dec_low. auto. +apply (VertexSet.diff_1 (VertexSet.inter_1 H3)). +apply (VertexSet.diff_1 (VertexSet.inter_2 H3)). +split. +apply move_merge_false; auto. case_eq (move_related g x); intros. +rewrite H4 in H1. inversion H1. reflexivity. +split. unfold g'. rewrite In_merge_vertex. split. +apply (in_interf_in x (fst_ext e) g). +apply (VertexSet.diff_1 (VertexSet.inter_1 H3)). +intro. generalize (proj2 (Aff_edge_aff e g HH HH0)). intro. +rewrite (compat_bool_move _ _ _ (Register.eq_sym _ _ H4)) in H5. +rewrite H5 in H1. inversion H1. +unfold g'. rewrite precolored_merge. intro. +elim (VertexSet.diff_2 (VertexSet.inter_1 H3)). apply (VertexSet.remove_3 H4). +unfold M in H0. +generalize (VertexSet.filter_1 (compat_move_up _ _) H0). intro. +generalize (VertexSet.filter_2 (compat_move_up _ _) H0). clear H0. intro. +assert (VertexSet.cardinal (preference_adj x g) = 1). +case_eq (eq_K 1 (VertexSet.cardinal (preference_adj x g))); intros. +generalize (eq_K_2 _ _ H2). clear H2. intro. auto. +rewrite H2 in H0. inversion H0. +generalize (eq_K_1 _ _ H2). intro. rewrite H3 in H0. simpl in H0. +split. +apply low_dec; auto. +intro. +destruct (VertexSet.union_1 H1). +elim (not_in_interf_self (snd_ext e) g). +rewrite H4 in H5. apply (VertexSet.diff_1 (VertexSet.inter_2 H5)). +elim (not_in_pref_self (snd_ext e) g). +rewrite H4 in H5. apply (VertexSet.diff_1 (VertexSet.inter_1 H5)). +split. +apply move_related_change_charac. +destruct (VertexSet.union_1 H1). +apply VertexSet.union_3. +apply VertexSet.inter_3. +apply (VertexSet.diff_1 (VertexSet.inter_2 H4)). +apply (VertexSet.diff_1 (VertexSet.inter_1 H4)). +apply VertexSet.union_2. +apply VertexSet.inter_3. +apply (VertexSet.diff_1 (VertexSet.inter_2 H4)). +apply (VertexSet.diff_1 (VertexSet.inter_1 H4)). +assumption. +split. +unfold g'. rewrite In_merge_vertex. split. +destruct (VertexSet.union_1 H1). +apply (in_interf_in x (snd_ext e) g). apply (VertexSet.diff_1 (VertexSet.inter_2 H4)). +apply (in_interf_in x (fst_ext e) g). apply (VertexSet.diff_1 (VertexSet.inter_2 H4)). +intro. destruct (VertexSet.union_1 H1). +elim (not_in_interf_self (snd_ext e) g). +rewrite H4 in H5. apply (VertexSet.diff_1 (VertexSet.inter_2 H5)). +elim (not_in_pref_self (snd_ext e) g). +rewrite H4 in H5. apply (VertexSet.diff_1 (VertexSet.inter_1 H5)). +unfold g'. rewrite precolored_merge. +intro. +destruct (VertexSet.union_1 H1). +elim (VertexSet.diff_2 (VertexSet.inter_1 H5)). apply (VertexSet.remove_3 H4). +elim (VertexSet.diff_2 (VertexSet.inter_1 H5)). apply (VertexSet.remove_3 H4). + +(* <= *) +destruct H. destruct H0. destruct H1. +assert (VertexSet.In x simplifyWL'). +unfold simplifyWL'. +case_eq (move_related g x); intros. +apply VertexSet.union_3. +unfold M. +apply VertexSet.filter_3. +apply compat_move_up. +assert (~Register.eq x (snd_ext e)) as Hsnd. +intro. elim (not_in_merge_snd e g HH HH0). rewrite H4 in H1. assumption. +generalize (move_merge_not_move _ _ _ _ HH0 n Hsnd H0 H3). intro. +destruct H4. +generalize (VertexSet.inter_1 H4). generalize (VertexSet.inter_2 H4). intros. +apply VertexSet.union_3. +apply VertexSet.inter_3. +apply VertexSet.diff_3. assumption. +unfold g' in H2. rewrite precolored_merge in H2. intro. elim H2. +apply VertexSet.remove_2. intro. elim (not_in_merge_snd _ _ HH HH0). +rewrite H8. assumption. +assumption. +apply VertexSet.diff_3. assumption. +unfold g' in H2. rewrite precolored_merge in H2. intro. elim H2. +apply VertexSet.remove_2. intro. elim (not_in_merge_snd _ _ HH HH0). +rewrite H8. assumption. +assumption. +generalize (VertexSet.inter_1 H4). generalize (VertexSet.inter_2 H4). intros. +apply VertexSet.union_2. +apply VertexSet.inter_3. +apply VertexSet.diff_3. assumption. +unfold g' in H2. +rewrite precolored_merge in H2. intro. elim H2. +apply VertexSet.remove_2. intro. elim (not_in_merge_snd _ _ HH HH0). +rewrite H8. assumption. +assumption. +apply VertexSet.diff_3. assumption. +unfold g' in H2. +rewrite precolored_merge in H2. intro. elim H2. +apply VertexSet.remove_2. intro. elim (not_in_merge_snd _ _ HH HH0). +rewrite H8. assumption. +assumption. + +assert (~Register.eq x (snd_ext e)) as Hsnd. +intro. rewrite H4 in H1. elim (not_in_merge_snd e g HH HH0 H1). +generalize (move_related_dec_1 _ _ _ HH HH0 n Hsnd H3 H0). intro. +unfold pref_degree in H4. rewrite (eq_K_1 _ _ H4). simpl. +case_eq (has_low_degree g k x); intros. +reflexivity. +generalize (move_merge_not_move _ _ _ _ HH0 n Hsnd H0 H3). intro. +assert (~Register.eq x (snd_ext e)). +intro. rewrite H7 in H1. elim (not_in_merge_snd e g HH HH0 H1). +generalize (merge_dec_interf _ _ _ _ HH HH0 H5 H n H7). intro. +destruct H8. destruct H6. +generalize (VertexSet.inter_2 H6). intro. +elim (interf_pref_conflict x (snd_ext e) g). +rewrite in_pref in H10. +split. +unfold Prefere. assumption. +unfold Interfere. rewrite <-in_interf. assumption. +elim (interf_pref_conflict x (fst_ext e) g). +generalize (VertexSet.inter_2 H6). intro. +rewrite in_pref in H10. +split. +unfold Prefere. assumption. +unfold Interfere. rewrite <-in_interf. auto. +apply VertexSet.union_2. +unfold simplifyWL_. +case_eq (has_low_degree g k x); intros. +apply VertexSet.union_2. + +WS_apply HWS. +split. assumption. +split. assumption. +split. unfold g' in H1. rewrite In_merge_vertex in H1. intuition. +unfold g' in H2. rewrite precolored_merge in H2. +intro. elim H2. apply VertexSet.remove_2. +intro. apply (not_in_merge_snd _ _ HH HH0). rewrite H6. assumption. +assumption. + +apply VertexSet.union_3. +rewrite HNMR. +apply VertexSet.filter_3. +apply compat_not_compat. apply compat_bool_move. +unfold N. apply VertexSet.filter_3. +apply eq_K_compat. +assert (~Register.eq x (snd_ext e)). +intro. rewrite H5 in H1. elim (not_in_merge_snd e g HH HH0 H1). +generalize (merge_dec_interf _ _ _ _ HH HH0 H4 H n H5). +intro. destruct H6. apply VertexSet.inter_3. +apply VertexSet.diff_3. assumption. +unfold g' in H2. rewrite precolored_merge in H2. +intro. elim H2. apply VertexSet.remove_2. auto. +assumption. +apply VertexSet.diff_3. assumption. +unfold g' in H2. rewrite precolored_merge in H2. +intro. elim H2. apply VertexSet.remove_2. auto. +assumption. +apply eq_K_1. +apply (merge_dec_K x e k g HH HH0 n). +intro. rewrite H5 in H1. elim (not_in_merge_snd e g HH HH0 H1). +assumption. +assumption. +rewrite H3. auto. + +unfold classify, get_spillWL, get_simplifyWL, get_movesWL, get_freezeWL; simpl. +case_eq (is_precolored (fst_ext e) g'); intros. +assumption. +case_eq (has_low_degree g' k (fst_ext e)); intros. +case_eq (move_related g' (fst_ext e)); intros. +assumption. +simpl. apply VertexSet.add_2; auto. +assumption. + +(* moves !!!!! *) +split; intros. +assert (EdgeSet.In e0 movesWL'). +unfold classify,get_spillWL, get_simplifyWL, get_movesWL, get_freezeWL in H; simpl in H; +destruct (is_precolored (fst_ext e) g'); +destruct (has_low_degree g' k (fst_ext e)); +destruct (move_related g' (fst_ext e)); simpl; assumption. +unfold movesWL' in H0. +unfold g'. rewrite <-AE_merge_wl. eassumption. +intros. rewrite <-In_graph_aff_edge_in_AE. +rewrite moves_AE. unfold moves. reflexivity. eassumption. + +destruct H. simpl. +assert (EdgeSet.In e0 movesWL'). +unfold movesWL'. unfold g', padj_fst_, padj_snd_. rewrite AE_merge_wl. +split; assumption. + +intros. rewrite <-In_graph_aff_edge_in_AE. +rewrite moves_AE. unfold moves. reflexivity. eassumption. +unfold classify,get_spillWL, get_simplifyWL, get_movesWL, get_freezeWL; simpl; +destruct (is_precolored (fst_ext e) g'); +destruct (has_low_degree g' k (fst_ext e)); +destruct (move_related g' (fst_ext e)); simpl; assumption. + +rewrite <-VertexSet.partition_2. rewrite Hmr. intuition. +apply compat_bool_move. +rewrite <-VertexSet.partition_1. rewrite Hmr. intuition. +apply compat_bool_move. +Qed. + +Lemma WS_coalesce : forall e ircg H H0, +WS_properties (merge e (irc_g ircg) H H0) + (irc_k ircg) + (merge_wl3 e ircg (merge e (irc_g ircg) H H0) H H0). + +Proof. +intros. rewrite <-(Hk ircg). apply ws_merge3. +Qed. \ No newline at end of file diff --git a/backend/MyAllocation.v b/backend/MyAllocation.v new file mode 100644 index 00000000..0d5f3e3d --- /dev/null +++ b/backend/MyAllocation.v @@ -0,0 +1,1434 @@ +Require Import IRC. +Require Import IRCColoring. +Require Import Graph_translation. +Require Import MyRegisters. +Require Import Locations. +Require Import RTLtyping. +Require Import ZArith. +Require Import AST. +Require Import Typed_interfgraphs. +Require Import Edges. +Require Import Graph_Facts. +Require Import Interference_adjacency. +Require Import InterfGraph. +Require Import Conventions. +Require Import Palettes. +Require Import InterfGraph_Construction. +Require Import WS. +Require Import Conservative_criteria. +Require Import IRC_graph. +Require Import IRC_Graph_Functions. +Require Import InterfGraphMapImp. +Require Import Registers. + +Import Props Edge RegFacts. + +Module ColFacts := FMapFacts.Facts ColorMap. + +Definition graph_coloring_aux x int_map float_map env := +match (Regs.get_type (Regs.P x) env) with +| Tint => match (map_to_coloring int_map (Regs.P x)) with + | Some (Regs.P z) => S (Local Z0 Tint) + | Some (Regs.M z) => R z + | None => S (Local (Zpos x) Tint) + end +| Tfloat => match (map_to_coloring float_map (Regs.P x)) with + | Some (Regs.P z) => S (Local Z0 Tfloat) + | Some (Regs.M z) => R z + | None => S (Local (Zpos x) Tfloat) + end +end. + +Definition reg_translation s := +Regset.fold (fun v s => VertexSet.add (Regs.reg_to_Reg v) s) s VertexSet.empty. + +Definition mreg_translation s := +MRegset.fold (fun v s => VertexSet.add (Regs.mreg_to_Reg v) s) s VertexSet.empty. + +Definition Typed_interfgraphs g env := +let regreg_interf_partition := +regreg_edge_type_partition (interf_reg_reg g) env in +let int_regreg_interf_edge := rr1 regreg_interf_partition in +let float_regreg_interf_edge := rr2 regreg_interf_partition in +let int_regreg_interf_reg := rr3 regreg_interf_partition in +let float_regreg_interf_reg := rr4 regreg_interf_partition in +let regmreg_interf_partition := +regmreg_edge_type_partition (interf_reg_mreg g) env in +let int_regmreg_interf_edge := rm1 regmreg_interf_partition in +let float_regmreg_interf_edge := rm2 regmreg_interf_partition in +let int_regmreg_interf_reg := rm3 regmreg_interf_partition in +let float_regmreg_interf_reg := rm4 regmreg_interf_partition in +let int_regmreg_interf_mreg := rm5 regmreg_interf_partition in +let float_regmreg_interf_mreg := rm6 regmreg_interf_partition in +let regreg_pref_partition := +regreg_edge_type_partition (pref_reg_reg g) env in +let int_regreg_pref_edge := rr1 regreg_pref_partition in +let float_regreg_pref_edge := rr2 regreg_pref_partition in +let int_regreg_pref_reg := rr3 regreg_pref_partition in +let float_regreg_pref_reg := rr4 regreg_pref_partition in +let regmreg_pref_partition := +regmreg_edge_type_partition (pref_reg_mreg g) env in +let int_regmreg_pref_edge := rm1 regmreg_pref_partition in +let float_regmreg_pref_edge := rm2 regmreg_pref_partition in +let int_regmreg_pref_reg := rm3 regmreg_pref_partition in +let float_regmreg_pref_reg := rm4 regmreg_pref_partition in +let int_regmreg_pref_mreg := rm5 regmreg_pref_partition in +let float_regmreg_pref_mreg := rm6 regmreg_pref_partition in +let int_regs := Regset.union int_regreg_interf_reg + (Regset.union int_regmreg_interf_reg + (Regset.union int_regreg_pref_reg int_regmreg_pref_reg)) in +let float_regs := Regset.union float_regreg_interf_reg + (Regset.union float_regmreg_interf_reg + (Regset.union float_regreg_pref_reg float_regmreg_pref_reg)) in +let int_mregs := MRegset.union int_regmreg_interf_mreg int_regmreg_pref_mreg in +let float_mregs := MRegset.union float_regmreg_interf_mreg float_regmreg_pref_mreg in +let int_Regs := VertexSet.union (reg_translation int_regs) (mreg_translation int_mregs) in +let float_Regs := VertexSet.union (reg_translation float_regs) (mreg_translation float_mregs) in +(int_Regs, +mkgraph int_regreg_interf_edge int_regmreg_interf_edge int_regreg_pref_edge int_regmreg_pref_edge, +float_Regs, +mkgraph float_regreg_interf_edge float_regmreg_interf_edge float_regreg_pref_edge float_regmreg_pref_edge). + +Lemma extremities_int_interf_graph : forall g env, +forall e, EdgeSet.In e (interfgraph_affinity_edges (snd (fst (fst (Typed_interfgraphs g env))))) \/ + EdgeSet.In e (interfgraph_interference_edges (snd (fst (fst (Typed_interfgraphs g env))))) -> + VertexSet.In (fst_ext e) (fst (fst (fst (Typed_interfgraphs g env)))) /\ + VertexSet.In (snd_ext e) (fst (fst (fst (Typed_interfgraphs g env)))). + +Proof. +Admitted. + +Lemma extremities_float_interf_graph : forall g env, +forall e, EdgeSet.In e (interfgraph_affinity_edges (snd (Typed_interfgraphs g env))) \/ + EdgeSet.In e (interfgraph_interference_edges (snd (Typed_interfgraphs g env))) -> + VertexSet.In (fst_ext e) (snd (fst (Typed_interfgraphs g env))) /\ + VertexSet.In (snd_ext e) (snd (fst (Typed_interfgraphs g env))). + +Proof. +Admitted. + +Definition my_graph_coloring g env := +let typed_graphs := Typed_interfgraphs g env in +let intR := fst (fst (fst typed_graphs)) in +let intG := snd (fst (fst typed_graphs)) in +let floatR := snd (fst typed_graphs) in +let floatG := snd typed_graphs in +let int_graph := graph_translation intG intR (extremities_int_interf_graph g env) in +let float_graph := graph_translation floatG floatR (extremities_float_interf_graph g env) in +let int_map := (IRC_map (graph_to_IRC_graph int_graph int_palette)) in +let float_map := (IRC_map (graph_to_IRC_graph float_graph float_palette)) in +fun x => graph_coloring_aux x int_map float_map env. + +Section Coloring_to_allocation. + +Variable g : graph. +Variable env : regenv. +Definition typed_graphs := Typed_interfgraphs g env. +Definition intR := fst (fst (fst typed_graphs)). +Definition intG := snd (fst (fst typed_graphs)). +Definition floatR := snd (fst typed_graphs). +Definition floatG := snd typed_graphs. +Definition int_graph := graph_translation intG intR (extremities_int_interf_graph g env). +Definition float_graph := graph_translation floatG floatR (extremities_float_interf_graph g env). +Definition int_map := (IRC_map (graph_to_IRC_graph int_graph int_palette)). +Definition float_map := (IRC_map (graph_to_IRC_graph float_graph float_palette)). +Definition int_coloring := map_to_coloring int_map. +Definition float_coloring := map_to_coloring float_map. + +Hypothesis temporaries_out : forall x, +In_graph (Regs.M x) int_graph -> ~List.In (R x) temporaries. + +Hypothesis correct_palette_int : forall x, +VertexSet.In x (precolored int_graph) -> VertexSet.In x int_palette. + +Hypothesis correct_palette_float : forall x, +VertexSet.In x (precolored float_graph) -> VertexSet.In x float_palette. + +Lemma proper_coloring_int : proper_coloring int_coloring int_graph int_palette. + +Proof. +intros. apply proper_coloring_IRC_aux. +intro. apply correct_palette_int. +Qed. + +Lemma proper_coloring_float : proper_coloring float_coloring float_graph float_palette. + +Proof. +intros. apply proper_coloring_IRC_aux. +intro. apply correct_palette_float. +Qed. + +Import SetoidList. + +Lemma exists_refl : forall x, +exists y, Regs.M x = Regs.M y. + +Proof. +intro x. exists x. auto. +Qed. + +Lemma mreg_int_palette : forall x, +VertexSet.In x int_palette -> +exists y, x = Regs.M y. + +Proof. +unfold int_palette. intros x H. +repeat (destruct (proj1 (Props.Dec.F.add_iff _ _ _) H); + [inversion H0;subst; apply exists_refl|generalize H0; clear H H0; intro H]). +elim (VertexSet.empty_1 H). +Qed. + +Lemma mreg_float_palette : forall x, +VertexSet.In x float_palette -> +exists y, x = Regs.M y. + +Proof. +unfold float_palette. intros x H. +repeat (destruct (proj1 (Props.Dec.F.add_iff _ _ _) H); + [inversion H0;subst; apply exists_refl|generalize H0; clear H H0; intro H]). +elim (VertexSet.empty_1 H). +Qed. + +Lemma register_heuristic_mreg : forall x r, +(IRC int_graph int_palette) x = Some r -> +exists y, r = Regs.M y. + +Proof. +intros x r H. +apply mreg_int_palette. +generalize (proper_coloring_IRC_aux int_graph int_palette correct_palette_int). +intro H0. +unfold proper_coloring in H0. +unfold proper_coloring_3 in H0. +do 2 destruct H0 as [_ H0]. +apply H0 with (x := x). +rewrite H. apply OptionReg.eq_refl. +Qed. + +Lemma register_heuristic_mreg_float : forall x r, +(IRC float_graph float_palette) x = Some r -> +exists y, r = Regs.M y. + +Proof. +intros x r H. +apply mreg_float_palette. +generalize (proper_coloring_IRC_aux float_graph float_palette correct_palette_float). +intro H0. +unfold proper_coloring in H0. +unfold proper_coloring_3 in H0. +do 2 destruct H0 as [_ H0]. +apply H0 with (x := x). +rewrite H. apply OptionReg.eq_refl. +Qed. + +Lemma int_palette_type : forall x, +VertexSet.In x int_palette -> +Regs.get_type x env = Tint. + +Proof. +unfold int_palette. intros x H. +repeat (destruct (proj1 (Props.Dec.F.add_iff _ _ _) H); + [inversion H0;subst; auto|generalize H0; clear H H0; intro H]). +elim (VertexSet.empty_1 H). +Qed. + +Lemma float_palette_type : forall x, +VertexSet.In x float_palette -> +Regs.get_type x env = Tfloat. + +Proof. +unfold float_palette. intros x H. +repeat (destruct (proj1 (Props.Dec.F.add_iff _ _ _) H); + [inversion H0;subst; auto|generalize H0; clear H H0; intro H]). +elim (VertexSet.empty_1 H). +Qed. + +Lemma register_heuristic_type_int : forall x r, +IRC int_graph int_palette x = Some r -> +Regs.get_type r env = Tint. + +Proof. +intros x r H. +apply int_palette_type. +generalize (proper_coloring_IRC_aux int_graph int_palette (correct_palette_int)). +intro H0. +unfold proper_coloring in H0. do 2 destruct H0 as [_ H0]. +unfold proper_coloring_3 in H0. +apply (H0 x r). +rewrite H. apply OptionReg.eq_refl. +Qed. + +Lemma register_heuristic_type_float : forall x r, +IRC float_graph float_palette x = Some r -> +Regs.get_type r env = Tfloat. + +Proof. +intros x r H. +apply float_palette_type. +generalize (proper_coloring_IRC_aux float_graph float_palette correct_palette_float). +intro H0. +unfold proper_coloring in H0. do 2 destruct H0 as [_ H0]. +unfold proper_coloring_3 in H0. +apply (H0 x r). +rewrite H. apply OptionReg.eq_refl. +Qed. + +Lemma Loc_reg_eq_type : forall x, +Regs.get_type (Regs.P x) env = Loc.type (my_graph_coloring g env x). + +Proof. +intro x. +unfold my_graph_coloring. +change (snd (fst (fst (Typed_interfgraphs g env)))) with intG. +change (fst (fst (fst (Typed_interfgraphs g env)))) with intR. +change (snd (fst (Typed_interfgraphs g env))) with floatR. +change (snd (Typed_interfgraphs g env)) with floatG. +fold int_graph; fold float_graph. +unfold graph_coloring_aux. +case_eq (Regs.get_type (Regs.P x) env); intros HH. +change (map_to_coloring (IRC_map (graph_to_IRC_graph int_graph int_palette))) with + (IRC int_graph int_palette). +case_eq (IRC int_graph int_palette (Regs.P x)); intros. +generalize (register_heuristic_mreg _ _ H). intro. destruct H0. rewrite H0. +generalize (register_heuristic_type_int _ _ H). +unfold Regs.get_type. rewrite H0. simpl. auto. +simpl. auto. +change (map_to_coloring (IRC_map (graph_to_IRC_graph float_graph float_palette))) with + (IRC float_graph float_palette). +case_eq (IRC float_graph float_palette (Regs.P x)); intros. +generalize (register_heuristic_mreg_float _ _ H). intro. destruct H0. rewrite H0. +generalize (register_heuristic_type_float _ _ H). +unfold Regs.get_type. rewrite H0. simpl. auto. +simpl. auto. +Qed. + +Lemma regreg_in_fst_partition : forall e s, +SetRegReg.In e s -> +env (fst e) = Tint -> +env (snd e) = Tint -> +SetRegReg.In e (rr1 (regreg_edge_type_partition s env)). + +Proof. +intros. +unfold regreg_edge_type_partition. +set (f:=(fun (e0 : SetRegReg.elt) (s0 : regregpartition) => + match env (fst e0) with + | Tint => + match env (snd e0) with + | Tint => + (SetRegReg.add e0 (rr1 s0), rr2 s0, + Regset.add (fst e0) (Regset.add (snd e0) (rr3 s0)), + rr4 s0) + | Tfloat => + (rr1 s0, rr2 s0, Regset.add (fst e0) (rr3 s0), + Regset.add (snd e0) (rr4 s0)) + end + | Tfloat => + match env (snd e0) with + | Tint => + (rr1 s0, rr2 s0, Regset.add (snd e0) (rr3 s0), + Regset.add (fst e0) (rr4 s0)) + | Tfloat => + (rr1 s0, SetRegReg.add e0 (rr2 s0), rr3 s0, + Regset.add (fst e0) (Regset.add (snd e0) (rr4 s0))) + end + end)). +unfold regregpartition in *. fold f. + +generalize (SetRegReg.empty,SetRegReg.empty, Regset.empty, Regset.empty). +generalize SetRegReg.elements_1. intro HH. +generalize (HH s e H). clear H HH. intro HH. +intro p. rewrite SetRegReg.fold_1. generalize p. clear p. +induction (SetRegReg.elements s). +inversion HH. +inversion HH. + subst. intro p. do 3 destruct p. simpl. + +assert (f a (t2,t3,t1,t0) = (SetRegReg.add a t2, t3, Regset.add (fst a) (Regset.add (snd a) t1),t0)). +unfold f. + +destruct H2. rewrite H in *. rewrite H2 in *. rewrite H0. rewrite H1. simpl. reflexivity. +rewrite H. +destruct H2. + +assert (forall x s1 s2 s3 s4, SetRegReg.In x s1 -> + SetRegReg.In x (rr1 (fold_left + (fun (a0 : SetRegReg.t*SetRegReg.t*Regset.t*Regset.t) (e0 : SetRegReg.elt) => f e0 a0) + l (s1, s2, s3, s4)))). + +clear H H0 H1 H2 HH IHl. + +induction l. +simpl. auto. +intros x s1 s2 s3 s4 H2. +simpl. + +assert (f a0 (s1,s2,s3,s4) = (SetRegReg.add a0 s1, s2,Regset.add (fst a0) (Regset.add (snd a0) s3), s4) \/ + f a0 (s1,s2,s3,s4) = (s1, SetRegReg.add a0 s2, s3, Regset.add (fst a0)(Regset.add (snd a0) s4)) \/ + f a0 (s1,s2,s3,s4) = (s1,s2,Regset.add (fst a0) s3, Regset.add (snd a0) s4)\/ + f a0 (s1,s2,s3,s4) = (s1,s2,Regset.add (snd a0) s3, Regset.add (fst a0) s4)). + +unfold f. +destruct (env (fst a0)); destruct (env (snd a0)); +unfold rr1,rr2,rr3,rr4; simpl; auto. + +destruct H. + +rewrite H. apply IHl. apply SetRegReg.add_2. assumption. + +destruct H. +rewrite H. +apply IHl. assumption. +destruct H; rewrite H. +apply IHl. assumption. +apply IHl. assumption. + +apply H4. apply SetRegReg.add_1. intuition. + +subst. simpl. intro p. apply IHl. assumption. +Qed. + +Lemma regreg_in_snd_partition : forall e s, +SetRegReg.In e s -> +env (fst e) = Tfloat -> +env (snd e) = Tfloat -> +SetRegReg.In e (rr2 (regreg_edge_type_partition s env)). + +Proof. +intros. +unfold regreg_edge_type_partition. +set (f:=(fun (e0 : SetRegReg.elt) (s0 : regregpartition) => + match env (fst e0) with + | Tint => + match env (snd e0) with + | Tint => + (SetRegReg.add e0 (rr1 s0), rr2 s0, + Regset.add (fst e0) (Regset.add (snd e0) (rr3 s0)), + rr4 s0) + | Tfloat => + (rr1 s0, rr2 s0, Regset.add (fst e0) (rr3 s0), + Regset.add (snd e0) (rr4 s0)) + end + | Tfloat => + match env (snd e0) with + | Tint => + (rr1 s0, rr2 s0, Regset.add (snd e0) (rr3 s0), + Regset.add (fst e0) (rr4 s0)) + | Tfloat => + (rr1 s0, SetRegReg.add e0 (rr2 s0), rr3 s0, + Regset.add (fst e0) (Regset.add (snd e0) (rr4 s0))) + end + end)). +unfold regregpartition in *. fold f. + +generalize (SetRegReg.empty,SetRegReg.empty, Regset.empty, Regset.empty). +generalize SetRegReg.elements_1. intro HH. +generalize (HH s e H). clear H HH. intro HH. +intro p. rewrite SetRegReg.fold_1. generalize p. clear p. +induction (SetRegReg.elements s). +inversion HH. +inversion HH. + subst. intro p. do 3 destruct p. simpl. + +assert (f a (t2,t3,t1,t0) = (t2, SetRegReg.add a t3, t1, Regset.add (fst a) (Regset.add (snd a) t0))). +unfold f. + +destruct H2. rewrite H in *. rewrite H2 in *. rewrite H0. rewrite H1. simpl. reflexivity. +rewrite H. +destruct H2. + +assert (forall x s1 s2 s3 s4, SetRegReg.In x s2 -> + SetRegReg.In x (rr2 (fold_left + (fun (a0 : SetRegReg.t*SetRegReg.t*Regset.t*Regset.t) (e0 : SetRegReg.elt) => f e0 a0) + l (s1, s2, s3, s4)))). + +clear H H0 H1 H2 HH IHl. + +induction l. +simpl. auto. +intros x s1 s2 s3 s4 H2. +simpl. + +assert (f a0 (s1,s2,s3,s4) = (SetRegReg.add a0 s1, s2,Regset.add (fst a0) (Regset.add (snd a0) s3), s4) \/ + f a0 (s1,s2,s3,s4) = (s1, SetRegReg.add a0 s2, s3, Regset.add (fst a0)(Regset.add (snd a0) s4)) \/ + f a0 (s1,s2,s3,s4) = (s1,s2,Regset.add (fst a0) s3, Regset.add (snd a0) s4)\/ + f a0 (s1,s2,s3,s4) = (s1,s2,Regset.add (snd a0) s3, Regset.add (fst a0) s4)). + +unfold f. +destruct (env (fst a0)); destruct (env (snd a0)); +unfold rr1,rr2,rr3,rr4; simpl; auto. + +destruct H. + +rewrite H. apply IHl. assumption. + +destruct H. +rewrite H. +apply IHl. apply SetRegReg.add_2. assumption. +destruct H; rewrite H. +apply IHl. assumption. +apply IHl. assumption. + +apply H4. apply SetRegReg.add_1. intuition. + +subst. simpl. intro p. apply IHl. assumption. +Qed. + +Lemma interf_int_regreg_translation : + interf_reg_reg (snd (fst (fst (Typed_interfgraphs g env)))) = + rr1 (regreg_edge_type_partition (interf_reg_reg g) env). + +Proof. +unfold Typed_interfgraphs. +case_eq (regreg_edge_type_partition (interf_reg_reg g) env); intros. +case_eq (regreg_edge_type_partition (pref_reg_reg g) env); intros. +case_eq (regmreg_edge_type_partition (interf_reg_mreg g) env); intros. +case_eq (regmreg_edge_type_partition (pref_reg_mreg g) env); intros. +simpl. reflexivity. +Qed. + +Lemma interf_float_regreg_translation : + interf_reg_reg (snd (Typed_interfgraphs g env)) = + rr2 (regreg_edge_type_partition (interf_reg_reg g) env). + +Proof. +unfold Typed_interfgraphs. +case_eq (regreg_edge_type_partition (interf_reg_reg g) env); intros. +case_eq (regreg_edge_type_partition (pref_reg_reg g) env); intros. +case_eq (regmreg_edge_type_partition (interf_reg_mreg g) env); intros. +case_eq (regmreg_edge_type_partition (pref_reg_mreg g) env); intros. +simpl. reflexivity. +Qed. + +Lemma correct_alloc_1 : check_coloring_1 g (my_graph_coloring g env) = true. + +Proof. +unfold check_coloring_1. +apply SetRegReg.for_all_1. + +unfold compat_bool. +intros x y H. destruct H as [H H0]. +rewrite H. rewrite H0. reflexivity. + +unfold SetRegReg.For_all. +intros x H. +generalize (Loc_reg_eq_type (fst x)). generalize (Loc_reg_eq_type (snd x)). +unfold my_graph_coloring in *. +change (snd (fst (fst (Typed_interfgraphs g env)))) with intG. +change (fst (fst (fst (Typed_interfgraphs g env)))) with intR. +change (snd (fst (Typed_interfgraphs g env))) with floatR. +change (snd (Typed_interfgraphs g env)) with floatG. +fold int_graph; fold float_graph. +unfold graph_coloring_aux. +change (map_to_coloring (IRC_map (graph_to_IRC_graph int_graph int_palette))) with + (IRC int_graph int_palette). +change (map_to_coloring (IRC_map (graph_to_IRC_graph float_graph float_palette))) with + (IRC float_graph float_palette). +intros Locty1 Locty2. +case_eq (Regs.get_type (Regs.P (fst x)) env); intros HH. +case_eq (Regs.get_type (Regs.P (snd x)) env); intros HH0. +case_eq (IRC int_graph int_palette (Regs.P (fst x))); intros. +destruct (register_heuristic_mreg (Regs.P (fst x)) t0 H0). rewrite H1. simpl. +case_eq (IRC int_graph int_palette (Regs.P (snd x))); intros. +destruct (register_heuristic_mreg (Regs.P (snd x)) t1 H2). rewrite H3. + +generalize (proper_coloring_IRC_aux int_graph int_palette (correct_palette_int)). +intro H4. unfold proper_coloring in H4. +destruct H4 as [H4 _]. +unfold proper_coloring_1 in H4. +assert (~Regs.eq (Regs.M x0) (Regs.M x1)). +apply (H4 (Regs.P (fst x), Regs.P (snd x), None)). +unfold Edge.interf_edge. auto. +unfold int_graph. +right. simpl. +apply regreg_IE_translation. unfold intG, typed_graphs. +rewrite interf_int_regreg_translation. +apply regreg_in_fst_partition. destruct x. auto. auto. auto. +change_rewrite. rewrite H0. rewrite H1. apply OptionReg.eq_refl. +change_rewrite. rewrite H2. rewrite H3. apply OptionReg.eq_refl. +destruct (Loc.eq (R x0) (R x1)). subst. +elim H5. inversion e. auto. +reflexivity. +destruct (Loc.eq (R x0) (S (Local (Zpos (snd x)) Tint))). +inversion e. +reflexivity. +case_eq (IRC int_graph int_palette (Regs.P (snd x))); intros. +destruct (register_heuristic_mreg (Regs.P (snd x)) t0 H1). rewrite H2. +destruct (Loc.eq (S (Local (Zpos (fst x)) Tint)) (R x0)). +inversion e. +reflexivity. +destruct (Loc.eq (S (Local (Zpos (fst x)) Tint)) (S (Local (Zpos (snd x)) Tint))). +inversion e. +elim (set_reg_reg_diff_ext _ _ (or_introl _ H) H3). +reflexivity. + +rewrite HH in *. rewrite HH0 in *. +set (l1 := match IRC int_graph int_palette (Regs.P (fst x)) with + | Some (Regs.P _) => S (Local 0 Tint) + | Some (Regs.M z) => R z + | None => S (Local (Zpos (fst x)) Tint) + end) in *. +set (l2 := match IRC float_graph float_palette (Regs.P (snd x)) with + | Some (Regs.P _) => S (Local 0 Tfloat) + | Some (Regs.M z) => R z + | None => S (Local (Zpos (snd x)) Tfloat) + end) in *. +destruct (Loc.eq l1 l2). rewrite e in Locty2. rewrite <-Locty1 in Locty2. congruence. +reflexivity. + +case_eq (Regs.get_type (Regs.P (snd x)) env); intros HH0. + +rewrite HH in *. rewrite HH0 in *. +set (l1 :=match IRC float_graph float_palette (Regs.P (fst x)) with + | Some (Regs.P _) => S (Local 0 Tfloat) + | Some (Regs.M z) => R z + | None => S (Local (Zpos (fst x)) Tfloat) + end ) in *. +set (l2 := match IRC int_graph int_palette (Regs.P (snd x)) with + | Some (Regs.P _) => S (Local 0 Tint) + | Some (Regs.M z) => R z + | None => S (Local (Zpos (snd x)) Tint) + end) in *. +destruct (Loc.eq l1 l2). rewrite e in Locty2. rewrite <-Locty1 in Locty2. congruence. +reflexivity. + +case_eq (IRC float_graph float_palette (Regs.P (fst x))); intros. +destruct (register_heuristic_mreg_float (Regs.P (fst x)) t0 H0). rewrite H1. simpl. +case_eq (IRC float_graph float_palette (Regs.P (snd x))); intros. +destruct (register_heuristic_mreg_float (Regs.P (snd x)) t1 H2). rewrite H3. + +generalize (proper_coloring_IRC_aux float_graph float_palette (correct_palette_float)). +intro H4. unfold proper_coloring in H4. +destruct H4 as [H4 _]. +unfold proper_coloring_1 in H4. +assert (~Regs.eq (Regs.M x0) (Regs.M x1)). +apply (H4 (Regs.P (fst x), Regs.P (snd x), None)). +unfold Edge.interf_edge. auto. +unfold float_graph. +right. simpl. +apply regreg_IE_translation. unfold floatG, typed_graphs. +rewrite interf_float_regreg_translation. +apply regreg_in_snd_partition. destruct x. auto. auto. auto. +change_rewrite. rewrite H0. rewrite H1. apply OptionReg.eq_refl. +change_rewrite. rewrite H2. rewrite H3. apply OptionReg.eq_refl. +destruct (Loc.eq (R x0) (R x1)). subst. +elim H5. inversion e. auto. +reflexivity. +destruct (Loc.eq (R x0) (S (Local (Zpos (snd x)) Tfloat))). +inversion e. +reflexivity. +case_eq (IRC float_graph float_palette (Regs.P (snd x))); intros. +destruct (register_heuristic_mreg_float (Regs.P (snd x)) t0 H1). rewrite H2. +destruct (Loc.eq (S (Local (Zpos (fst x)) Tfloat)) (R x0)). +inversion e. +reflexivity. +destruct (Loc.eq (S (Local (Zpos (fst x)) Tfloat)) (S (Local (Zpos (snd x)) Tfloat))). +inversion e. +elim (set_reg_reg_diff_ext _ _ (or_introl _ H) H3). +reflexivity. +Qed. + +Lemma interf_int_regmreg_translation : + interf_reg_mreg (snd (fst (fst (Typed_interfgraphs g env)))) = + rm1 (regmreg_edge_type_partition (interf_reg_mreg g) env). + +Proof. +unfold Typed_interfgraphs. +case_eq (regreg_edge_type_partition (interf_reg_reg g) env); intros. +case_eq (regreg_edge_type_partition (pref_reg_reg g) env); intros. +case_eq (regmreg_edge_type_partition (interf_reg_mreg g) env); intros. +case_eq (regmreg_edge_type_partition (pref_reg_mreg g) env); intros. +simpl. reflexivity. +Qed. + +Lemma interf_float_regmreg_translation : + interf_reg_mreg (snd (Typed_interfgraphs g env)) = + rm2 (regmreg_edge_type_partition (interf_reg_mreg g) env). + +Proof. +unfold Typed_interfgraphs. +case_eq (regreg_edge_type_partition (interf_reg_reg g) env); intros. +case_eq (regreg_edge_type_partition (pref_reg_reg g) env); intros. +case_eq (regmreg_edge_type_partition (interf_reg_mreg g) env); intros. +case_eq (regmreg_edge_type_partition (pref_reg_mreg g) env); intros. +simpl. reflexivity. +Qed. + +Lemma regmreg_in_fst_partition : forall e s, +SetRegMreg.In e s -> +env (fst e) = Tint -> +mreg_type (snd e) = Tint -> +SetRegMreg.In e (rm1 (regmreg_edge_type_partition s env)). + +Proof. +intros. +unfold regmreg_edge_type_partition. +set (f := (fun (e0 : SetRegMreg.elt) (s0 : regmregpartition) => + match env (fst e0) with + | Tint => + match mreg_type (snd e0) with + | Tint => + (SetRegMreg.add e0 (rm1 s0), rm2 s0, + Regset.add (fst e0) (rm3 s0), rm4 s0, + MRegset.add (snd e0) (rm5 s0), rm6 s0) + | Tfloat => + (rm1 s0, rm2 s0, Regset.add (fst e0) (rm3 s0), rm4 s0, + rm5 s0, MRegset.add (snd e0) (rm6 s0)) + end + | Tfloat => + match mreg_type (snd e0) with + | Tint => + (rm1 s0, rm2 s0, rm3 s0, Regset.add (fst e0) (rm4 s0), + MRegset.add (snd e0) (rm5 s0), rm6 s0) + | Tfloat => + (rm1 s0, SetRegMreg.add e0 (rm2 s0), rm3 s0, + Regset.add (fst e0) (rm4 s0), rm5 s0, + MRegset.add (snd e0) (rm6 s0)) + end + end)). +unfold regmregpartition in *. fold f. + +generalize (SetRegMreg.empty, SetRegMreg.empty, Regset.empty, Regset.empty, MRegset.empty, MRegset.empty). +generalize SetRegMreg.elements_1. intro HH. +generalize (HH s e H). clear H HH. intro HH. +intro p. rewrite SetRegMreg.fold_1. generalize p. clear p. +induction (SetRegMreg.elements s). +inversion HH. +inversion HH. + subst. intro p. do 5 destruct p. simpl. + +assert (f a (t4,t5,t3,t2,t1,t0) = (SetRegMreg.add a t4, t5, Regset.add (fst a) t3, t2, MRegset.add (snd a) t1, t0)). +unfold f. + +destruct H2. rewrite H in *. rewrite H2 in *. rewrite H0. rewrite H1. simpl. reflexivity. +rewrite H. +destruct H2. + +assert (forall x s1 s2 s3 s4 s5 s6, SetRegMreg.In x s1 -> + SetRegMreg.In x (rm1 (fold_left + (fun (a0 : SetRegMreg.t * SetRegMreg.t * Regset.t * Regset.t *MRegset.t * MRegset.t) + (e0 : SetRegMreg.elt) => f e0 a0) + l (s1, s2, s3, s4, s5, s6)))). + +clear H H0 H1 H2 HH IHl. + +induction l. +simpl. auto. +intros x s1 s2 s3 s4 s5 s6 H2. +simpl. + +assert (f a0 (s1,s2, s3, s4, s5, s6) = (SetRegMreg.add a0 s1, s2, Regset.add (fst a0) s3, s4, MRegset.add (snd a0) s5, s6) \/ + f a0 (s1,s2,s3,s4,s5,s6) = (s1, SetRegMreg.add a0 s2, s3, Regset.add (fst a0) s4, s5, MRegset.add (snd a0) s6) \/ + f a0 (s1,s2,s3,s4,s5,s6) = (s1,s2,Regset.add (fst a0) s3, s4,s5,MRegset.add (snd a0) s6) \/ + f a0 (s1,s2,s3,s4,s5,s6) = (s1,s2,s3, Regset.add (fst a0) s4, MRegset.add (snd a0) s5, s6)). + +unfold f. +destruct (env (fst a0)); destruct (mreg_type (snd a0)); auto. + +destruct H. + +rewrite H. apply IHl. apply SetRegMreg.add_2. assumption. + +destruct H. +rewrite H. +apply IHl. assumption. +destruct H; rewrite H; apply IHl; assumption. + +apply H4. apply SetRegMreg.add_1. intuition. + +subst. simpl. intro p. apply IHl. assumption. +Qed. + +Lemma regmreg_in_snd_partition : forall e s, +SetRegMreg.In e s -> +env (fst e) = Tfloat -> +mreg_type (snd e) = Tfloat -> +SetRegMreg.In e (rm2 (regmreg_edge_type_partition s env)). + +Proof. +intros. +unfold regmreg_edge_type_partition. +set (f := (fun (e0 : SetRegMreg.elt) (s0 : regmregpartition) => + match env (fst e0) with + | Tint => + match mreg_type (snd e0) with + | Tint => + (SetRegMreg.add e0 (rm1 s0), rm2 s0, + Regset.add (fst e0) (rm3 s0), rm4 s0, + MRegset.add (snd e0) (rm5 s0), rm6 s0) + | Tfloat => + (rm1 s0, rm2 s0, Regset.add (fst e0) (rm3 s0), rm4 s0, + rm5 s0, MRegset.add (snd e0) (rm6 s0)) + end + | Tfloat => + match mreg_type (snd e0) with + | Tint => + (rm1 s0, rm2 s0, rm3 s0, Regset.add (fst e0) (rm4 s0), + MRegset.add (snd e0) (rm5 s0), rm6 s0) + | Tfloat => + (rm1 s0, SetRegMreg.add e0 (rm2 s0), rm3 s0, + Regset.add (fst e0) (rm4 s0), rm5 s0, + MRegset.add (snd e0) (rm6 s0)) + end + end)). +unfold regmregpartition in *. fold f. + +generalize (SetRegMreg.empty, SetRegMreg.empty, Regset.empty, Regset.empty, MRegset.empty, MRegset.empty). +generalize SetRegMreg.elements_1. intro HH. +generalize (HH s e H). clear H HH. intro HH. +intro p. rewrite SetRegMreg.fold_1. generalize p. clear p. +induction (SetRegMreg.elements s). +inversion HH. +inversion HH. + subst. intro p. do 5 destruct p. simpl. + +assert (f a (t4,t5,t3,t2,t1,t0) = (t4, SetRegMreg.add a t5, t3, Regset.add (fst a) t2, t1, MRegset.add (snd a) t0)). +unfold f. + +destruct H2. rewrite H in *. rewrite H2 in *. rewrite H0. rewrite H1. simpl. reflexivity. +rewrite H. +destruct H2. + +assert (forall x s1 s2 s3 s4 s5 s6, SetRegMreg.In x s2 -> + SetRegMreg.In x (rm2 (fold_left + (fun (a0 : SetRegMreg.t * SetRegMreg.t * Regset.t * Regset.t *MRegset.t * MRegset.t) + (e0 : SetRegMreg.elt) => f e0 a0) + l (s1, s2, s3, s4, s5, s6)))). + +clear H H0 H1 H2 HH IHl. + +induction l. +simpl. auto. +intros x s1 s2 s3 s4 s5 s6 H2. +simpl. + +assert (f a0 (s1,s2, s3, s4, s5, s6) = (SetRegMreg.add a0 s1, s2, Regset.add (fst a0) s3, s4, MRegset.add (snd a0) s5, s6) \/ + f a0 (s1,s2,s3,s4,s5,s6) = (s1, SetRegMreg.add a0 s2, s3, Regset.add (fst a0) s4, s5, MRegset.add (snd a0) s6) \/ + f a0 (s1,s2,s3,s4,s5,s6) = (s1,s2,Regset.add (fst a0) s3, s4,s5,MRegset.add (snd a0) s6) \/ + f a0 (s1,s2,s3,s4,s5,s6) = (s1,s2,s3, Regset.add (fst a0) s4, MRegset.add (snd a0) s5, s6)). + +unfold f. +destruct (env (fst a0)); destruct (mreg_type (snd a0)); auto. + +destruct H. + +rewrite H. apply IHl. assumption. + +destruct H. +rewrite H. +apply IHl. apply SetRegMreg.add_2. assumption. +destruct H; rewrite H; apply IHl; assumption. + +apply H4. apply SetRegMreg.add_1. intuition. + +subst. simpl. intro p. apply IHl. assumption. +Qed. + +Section fold_assoc_map. + +Variable A : Type. + +Lemma fold_left_compat_map : forall (f : ColorMap.t Regs.t -> A -> ColorMap.t Regs.t) l e e', +ColorMap.Equal e e' -> +(forall e1 e2 a, ColorMap.Equal e1 e2 -> ColorMap.Equal (f e1 a) (f e2 a)) -> +ColorMap.Equal (fold_left f l e) (fold_left f l e'). + +Proof. +intro f;induction l;simpl. +auto. +intros e e' H H0 H1. +apply (IHl (f e a) (f e' a)). +apply H0;assumption. +assumption. +Qed. + +Lemma fold_left_assoc_map : forall l (f : ColorMap.t Regs.t -> A -> ColorMap.t Regs.t) x h, +(forall (y z : A) s, ColorMap.Equal (f (f s y) z) (f (f s z) y)) -> +(forall e1 e2 a, ColorMap.Equal e1 e2 -> ColorMap.Equal (f e1 a) (f e2 a)) -> +ColorMap.Equal (fold_left f (h :: l) x) (f (fold_left f l x) h). + +Proof. +induction l;simpl;intros f x h H H0. +intuition. +rewrite <-IHl;simpl;try assumption. +apply fold_left_compat_map;[apply H|];auto. +Qed. + +End fold_assoc_map. + +Lemma mreg_refl_coloring_aux : forall x gpalette, +VertexSet.In x (precolored (irc_g gpalette)) -> +VertexSet.Subset (precolored (irc_g gpalette)) (pal gpalette) -> +OptionReg.eq (map_to_coloring (IRC_map gpalette) x) (Some x). + +Proof. +intros. functional induction IRC_map gpalette; simpl in *. + +(* simplify *) +generalize (simplify_inv _ _ e). intro. +generalize (simplify_inv2 _ _ e). intro. destruct H2. simpl in *. clear e. +rewrite H2 in *. clear H2. unfold available_coloring. +set (palette := pal g0) in *. set (wl := irc_wl g0) in *. set (g1 := irc_g g0) in *. +case_eq ( VertexSet.choose + (VertexSet.diff palette + (forbidden_colors r + (IRC_map + (simplify_irc r g0 + (VertexSet.choose_1 (s:=get_simplifyWL wl) x0))) g1))). +intros. unfold map_to_coloring. +rewrite ColFacts.add_neq_o. +apply IHt0. unfold simplify_irc. simpl. +rewrite precolored_remove_vertex. apply VertexSet.remove_2. +intro. rewrite <-H3 in H. +generalize (In_simplify_props _ _ _ _ _ _ _ _ (VertexSet.choose_1 H1) (refl_equal _) (HWS_irc g0)). intro. +destruct H4. destruct H5. destruct H6. elim H7. auto. auto. +unfold simplify_irc. simpl. rewrite precolored_remove_vertex. +intro. intro. apply H0. apply (VertexSet.remove_3 H3). +intro. rewrite <-H3 in H. +generalize (In_simplify_props _ _ _ _ _ _ _ _ (VertexSet.choose_1 H1) (refl_equal _) (HWS_irc g0)). intro. +destruct H4. destruct H5. destruct H6. elim H7. auto. intro. +apply IHt0. unfold simplify_irc. simpl. +rewrite precolored_remove_vertex. apply VertexSet.remove_2. +intro. rewrite <-H3 in H. +generalize (In_simplify_props _ _ _ _ _ _ _ _ (VertexSet.choose_1 H1) (refl_equal _) (HWS_irc g0)). intro. +destruct H4. destruct H5. destruct H6. elim H7. auto. auto. +unfold simplify_irc. simpl. rewrite precolored_remove_vertex. +intro. intro. apply H0. apply (VertexSet.remove_3 H3). + +(* coalesce *) +assert (forall e', EdgeSet.In e' (get_movesWL (irc_wl g0)) -> In_graph_edge e' (irc_g g0)). +intros. +generalize (In_move_props _ _ _ _ _ _ _ _ H1 (refl_equal _) (HWS_irc g0)). +intuition. +generalize (coalesce_inv _ _ e0). simpl. intro. +generalize (coalesce_inv_2 _ _ e0). intro. destruct H3. destruct H3. simpl in H3. rewrite H3 in *. clear H3. +generalize (any_coalescible_edge_1 _ _ _ _ H1 H2). +intro. destruct H3. +rewrite <-(moves_AE _ _ _ (HWS_irc g0)) in H4. +generalize (proj2 (proj1 (In_graph_aff_edge_in_AE _ _) H4)). intro. +generalize (any_coalescible_edge_2 _ _ _ _ H1 H2). intro. +unfold complete_coloring. +case_eq (ColorMap.find (elt:=Regs.t) (fst_ext e1) + (IRC_map (merge_irc e1 g0 x0 x1))). +intros. +unfold map_to_coloring. +rewrite ColFacts.add_neq_o. +apply IHt0. +assert (Edge.aff_edge e1). +rewrite (moves_AE _ _ _ (HWS_irc g0)) in H4. +generalize (In_move_props _ _ _ _ _ _ _ _ H4 (refl_equal _) (HWS_irc g0)). +intuition. +unfold merge_irc. simpl. +rewrite (precolored_merge _ _ H5 H8 _). +apply VertexSet.remove_2. intro. rewrite H9 in H6. elim H6. auto. auto. +unfold VertexSet.Subset in *. +intros. unfold merge_irc in *. simpl in *. +apply H0. rewrite precolored_merge in H8. apply (VertexSet.remove_3 H8). +intro. rewrite H8 in H6. elim H6. auto. +intro. +apply IHt0. +assert (Edge.aff_edge e1). +rewrite (moves_AE _ _ _ (HWS_irc g0)) in H4. +generalize (In_move_props _ _ _ _ _ _ _ _ H4 (refl_equal _) (HWS_irc g0)). +intuition. +unfold merge_irc. simpl. +rewrite (precolored_merge _ _ H5 H8 _). +apply VertexSet.remove_2. intro. rewrite H9 in H6. elim H6. auto. auto. +unfold VertexSet.Subset in *. +intros. unfold merge_irc in *. simpl in *. +apply H0. rewrite precolored_merge in H8. apply (VertexSet.remove_3 H8). + +(* freeze *) +generalize (freeze_inv _ _ e1). intro. +generalize (freeze_inv2 _ _ e1). intro. destruct H2. destruct H2. simpl in *. clear e1. +rewrite H2 in *. clear H2. unfold delete_preference_edges_irc2 in *. simpl in *. +apply IHt0. +rewrite precolored_delete_preference_edges. assumption. +unfold VertexSet.Subset in *. +intros. rewrite precolored_delete_preference_edges in H2. auto. + +(* spill *) +generalize e2. clear e e0 e1 e2. intro e. +generalize (spill_inv _ _ e). intro. +generalize (spill_inv2 _ _ e). intro. destruct H2. simpl in *. clear e. +rewrite H2 in *. clear H2. unfold available_coloring. +set (palette := pal g0) in *. set (wl := irc_wl g0) in *. set (g1 := irc_g g0) in *. +case_eq ( VertexSet.choose + (VertexSet.diff palette + (forbidden_colors r + (IRC_map + (spill_irc r g0 + (lowest_cost_in r (get_spillWL wl) g1 x0))) g1))). +intros. unfold map_to_coloring. +rewrite ColFacts.add_neq_o. +apply IHt0. unfold spill_irc. simpl. +rewrite precolored_remove_vertex. apply VertexSet.remove_2. +intro. rewrite <-H3 in H. +generalize (In_spill_props _ _ _ _ _ _ _ _ (lowest_cost_in _ _ _ H1) (refl_equal _) (HWS_irc g0)). intro. +destruct H4. destruct H5. elim H6. auto. auto. +unfold spill_irc. simpl. rewrite precolored_remove_vertex. +intro. intro. apply H0. apply (VertexSet.remove_3 H3). +intro. rewrite <-H3 in H. +generalize (In_spill_props _ _ _ _ _ _ _ _ (lowest_cost_in _ _ _ H1) (refl_equal _) (HWS_irc g0)). intro. +destruct H4. destruct H5. elim H6. auto. intro. +apply IHt0. unfold spill_irc. simpl. +rewrite precolored_remove_vertex. apply VertexSet.remove_2. +intro. rewrite <-H3 in H. +generalize (In_spill_props _ _ _ _ _ _ _ _ (lowest_cost_in _ _ _ H1) (refl_equal _) (HWS_irc g0)). intro. +destruct H4. destruct H5. elim H6. auto. auto. +unfold simplify_irc. simpl. rewrite precolored_remove_vertex. +intro. intro. apply H0. apply (VertexSet.remove_3 H3). + +(* ending case *) +set (palette := pal g0) in *. +set (g1 := irc_g g0) in *. +assert (map_to_coloring (precoloring_map g1) x = Some x). +unfold precoloring_map. +rewrite VertexSet.fold_1. +generalize VertexSet.elements_1. intro HH. +generalize (HH (precolored g1) x H). clear HH. intro HH. +generalize (NoDupA_elements (precolored g1)). intro HHH. +induction (VertexSet.elements (precolored g1)). +simpl. inversion HH. + +unfold map_to_coloring. +rewrite fold_left_assoc_map. +inversion HH. subst. +rewrite ColFacts.add_eq_o. +inversion H2; subst; auto. +apply Regs.eq_sym. auto. +subst. +unfold map_to_coloring in IHl. unfold Regs.t. +rewrite ColFacts.add_neq_o. +apply IHl. assumption. inversion HHH. assumption. +inversion HHH. subst. intro H6. +elim H4. +inversion H6; subst; auto. + +intros. +unfold ColorMap.Equal. +intro. +destruct (Regs.eq_dec z y0). +rewrite ColFacts.add_eq_o. +destruct (Regs.eq_dec y y0). +rewrite ColFacts.add_eq_o. +inversion e3; inversion e4; subst; rewrite H6; auto. +auto. +rewrite ColFacts.add_neq_o. +rewrite ColFacts.add_eq_o. +reflexivity. +assumption. +assumption. +assumption. +rewrite ColFacts.add_neq_o. +destruct (Regs.eq_dec y y0). +rewrite ColFacts.add_eq_o. +rewrite ColFacts.add_eq_o. +reflexivity. +assumption. +assumption. +rewrite ColFacts.add_neq_o. +rewrite ColFacts.add_neq_o. +rewrite ColFacts.add_neq_o. +reflexivity. +assumption. +assumption. +assumption. +assumption. + +intros. +apply ColFacts.add_m. +apply Regs.eq_refl. +reflexivity. +assumption. + +rewrite H1. apply OptionReg.eq_refl. +Qed. + +Lemma mreg_refl_coloring : forall x g palette, +VertexSet.In x (precolored g) -> +VertexSet.Subset (precolored g) palette -> +OptionReg.eq (IRC g palette x) (Some x). + +Proof. +intros. +apply mreg_refl_coloring_aux; +unfold graph_to_IRC_graph; simpl; auto. +Qed. + +Lemma loc_type_reg_type_equiv : forall x, +Loc.type (R x) = Regs.get_type (Regs.M x) env. + +Proof. +intro x. +unfold Loc.type. unfold Regs.get_type. reflexivity. +Qed. + +Lemma correct_alloc_2 : check_coloring_2 g (my_graph_coloring g env) = true. + +Proof. +unfold check_coloring_2. +apply SetRegMreg.for_all_1. + +unfold compat_bool. +intros x y H. destruct H as [H H0]. +rewrite H. rewrite H0. reflexivity. + +unfold SetRegMreg.For_all. +intros x H. +generalize (Loc_reg_eq_type (fst x)). generalize (loc_type_reg_type_equiv (snd x)). +unfold my_graph_coloring in *. +change (snd (fst (fst (Typed_interfgraphs g env)))) with intG. +change (fst (fst (fst (Typed_interfgraphs g env)))) with intR. +change (snd (fst (Typed_interfgraphs g env))) with floatR. +change (snd (Typed_interfgraphs g env)) with floatG. +fold int_graph; fold float_graph. +unfold graph_coloring_aux. +change (map_to_coloring (IRC_map (graph_to_IRC_graph int_graph int_palette))) with + (IRC int_graph int_palette). +change (map_to_coloring (IRC_map (graph_to_IRC_graph float_graph float_palette))) with + (IRC float_graph float_palette). +intros Locty1 Locty2. +case_eq (Regs.get_type (Regs.P (fst x)) env); intros HH. rewrite HH in *. +case_eq (Regs.get_type (Regs.M (snd x)) env); intros HH0. +case_eq (IRC int_graph int_palette (Regs.P (fst x))); intros. rewrite H0 in *. +destruct (register_heuristic_mreg (Regs.P (fst x)) t0 H0). rewrite H1 in *. simpl. +case_eq (IRC int_graph int_palette (Regs.M (snd x))); intros. +destruct (register_heuristic_mreg (Regs.M (snd x)) t1 H2). + +generalize (proper_coloring_IRC_aux int_graph int_palette (correct_palette_int)). +intro H4. unfold proper_coloring in H4. +destruct H4 as [H4 _]. +unfold proper_coloring_1 in H4. +assert (~Regs.eq (Regs.M x0) (Regs.M x1)). +apply (H4 (Regs.P (fst x), Regs.M (snd x), None)). +unfold Edge.interf_edge. auto. +unfold int_graph. +right. simpl. +apply regmreg_IE_translation. unfold intG, typed_graphs. +rewrite interf_int_regmreg_translation. +apply regmreg_in_fst_partition. destruct x. auto. auto. auto. +change_rewrite. rewrite H0. apply OptionReg.eq_refl. +change_rewrite. rewrite H2. rewrite H3. apply OptionReg.eq_refl. +destruct (Loc.eq (R x0) (R x1)). subst. +elim H5. inversion e. auto. + +destruct (Loc.eq (R x0) (R (snd x))). inversion e. clear e. +generalize (proper_coloring_IRC_aux int_graph int_palette correct_palette_int). +intro H6. unfold proper_coloring in H6. +destruct H6 as [H6 HH5]. destruct HH5 as [HH5 _]. +unfold proper_coloring_1 in H6. +assert (~Regs.req t0 (Regs.M (snd x))). +apply (H6 (Regs.P (fst x), Regs.M (snd x), None)). +unfold Edge.interf_edge. auto. +unfold int_graph. unfold intG, typed_graphs. +right. simpl. +apply regmreg_IE_translation. simpl. +apply regmreg_in_fst_partition. destruct x. auto. +auto. +auto. +change_rewrite. rewrite H0. rewrite H1. apply OptionReg.eq_refl. +change_rewrite. apply mreg_refl_coloring. subst. +apply (proj2 (precolored_equiv _ _)). +unfold is_precolored. simpl. +split. auto. + +assert (EdgeSet.In (Regs.reg_to_Reg (fst x), Regs.mreg_to_Reg (snd x), None) + (IE int_graph)). +unfold int_graph. +apply regmreg_IE_translation. destruct x. simpl. +case_eq (regreg_edge_type_partition (interf_reg_reg g) env); intros. +case_eq (regreg_edge_type_partition (pref_reg_reg g) env); intros. +case_eq (regmreg_edge_type_partition (interf_reg_mreg g) env); intros. +case_eq (regmreg_edge_type_partition (pref_reg_mreg g) env); intros. +simpl. +replace p2 with (fst (regmreg_edge_type_partition (interf_reg_mreg g) env)). +apply regmreg_in_fst_partition. auto. +assumption. +assumption. +rewrite H7. auto. +apply (proj2 (In_graph_edge_in_ext (Regs.P (fst x), Regs.M (snd x), None) + _ (or_intror _ H1))). +assumption. + +subst. elim H8. apply Regs.eq_refl. reflexivity. + +assert (OptionReg.eq (IRC int_graph int_palette (Regs.M (snd x))) (Some (Regs.M (snd x)))) as Hsnd. +apply mreg_refl_coloring. subst. +apply (proj2 (precolored_equiv _ _)). +unfold is_precolored. simpl. +split. auto. + +assert (EdgeSet.In (Regs.reg_to_Reg (fst x), Regs.mreg_to_Reg (snd x), None) + (IE int_graph)). +unfold int_graph. +apply regmreg_IE_translation. destruct x. simpl. +case_eq (regreg_edge_type_partition (interf_reg_reg g) env); intros. +case_eq (regreg_edge_type_partition (pref_reg_reg g) env); intros. +case_eq (regmreg_edge_type_partition (interf_reg_mreg g) env); intros. +case_eq (regmreg_edge_type_partition (pref_reg_mreg g) env); intros. +simpl. +replace p2 with (fst (regmreg_edge_type_partition (interf_reg_mreg g) env)). +apply regmreg_in_fst_partition. auto. +assumption. +assumption. +rewrite H4. auto. +apply (proj2 (In_graph_edge_in_ext (Regs.P (fst x), Regs.M (snd x), None) + _ (or_intror _ H1))). +assumption. + +rewrite H2 in Hsnd. inversion Hsnd. + +destruct (Loc.eq (S (Local (Zpos (fst x)) Tint)) (R (snd x))). +inversion e. reflexivity. + +case_eq (IRC int_graph int_palette (Regs.P (fst x))); intros. rewrite H0 in Locty2. +case_eq t0; intros; rewrite H1 in *. +destruct (Loc.eq (S (Local 0 Tint)) (R (snd x))). +inversion e. reflexivity. +destruct (Loc.eq (R m) (R (snd x))). +inversion e. subst. unfold Regs.get_type in HH0. unfold Loc.type in Locty2. congruence. +reflexivity. +destruct (Loc.eq (S (Local (Zpos (fst x)) Tint)) (R (snd x))). +inversion e. reflexivity. + +case_eq (Regs.get_type (Regs.M (snd x)) env); intros HH0. + +case_eq (IRC float_graph float_palette (Regs.P (fst x))); intros. rewrite H0 in Locty2. +case_eq t0; intros; rewrite H1 in *. +destruct (Loc.eq (S (Local 0 Tfloat)) (R (snd x))). +inversion e. reflexivity. +destruct (Loc.eq (R m) (R (snd x))). +inversion e. subst. rewrite HH in Locty2. unfold Regs.get_type in HH0. unfold Loc.type in Locty2. congruence. +reflexivity. +destruct (Loc.eq (S (Local (Zpos (fst x)) Tfloat)) (R (snd x))). +inversion e. reflexivity. + +case_eq (IRC float_graph float_palette (Regs.P (fst x))); intros. rewrite H0 in *. +destruct (register_heuristic_mreg_float (Regs.P (fst x)) t0 H0). rewrite H1 in *. simpl. +case_eq (IRC float_graph float_palette (Regs.M (snd x))); intros. +destruct (register_heuristic_mreg_float (Regs.M (snd x)) t1 H2). + +generalize (proper_coloring_IRC_aux float_graph float_palette (correct_palette_float)). +intro H4. unfold proper_coloring in H4. +destruct H4 as [H4 _]. +unfold proper_coloring_1 in H4. +assert (~Regs.eq (Regs.M x0) (Regs.M x1)). +apply (H4 (Regs.P (fst x), Regs.M (snd x), None)). +unfold Edge.interf_edge. auto. +unfold float_graph. +right. simpl. +apply regmreg_IE_translation. unfold floatG, typed_graphs. +rewrite interf_float_regmreg_translation. +apply regmreg_in_snd_partition. destruct x. auto. auto. auto. +change_rewrite. rewrite H0. apply OptionReg.eq_refl. +change_rewrite. rewrite H2. rewrite H3. apply OptionReg.eq_refl. +destruct (Loc.eq (R x0) (R x1)). subst. +elim H5. inversion e. auto. + +destruct (Loc.eq (R x0) (R (snd x))). inversion e. clear e. +generalize (proper_coloring_IRC_aux float_graph float_palette correct_palette_float). +intro H6. unfold proper_coloring in H6. +destruct H6 as [H6 HH5]. destruct HH5 as [HH5 _]. +unfold proper_coloring_1 in H6. +assert (~Regs.req t0 (Regs.M (snd x))). +apply (H6 (Regs.P (fst x), Regs.M (snd x), None)). +unfold Edge.interf_edge. auto. +unfold int_graph. unfold floatG, typed_graphs. +right. simpl. +apply regmreg_IE_translation. simpl. +apply regmreg_in_snd_partition. destruct x. auto. +auto. +auto. +change_rewrite. rewrite H0. rewrite H1. apply OptionReg.eq_refl. +change_rewrite. apply mreg_refl_coloring. subst. +apply (proj2 (precolored_equiv _ _)). +unfold is_precolored. simpl. +split. auto. + +assert (EdgeSet.In (Regs.reg_to_Reg (fst x), Regs.mreg_to_Reg (snd x), None) + (IE float_graph)). +unfold int_graph. +apply regmreg_IE_translation. destruct x. simpl. +case_eq (regreg_edge_type_partition (interf_reg_reg g) env); intros. +case_eq (regreg_edge_type_partition (pref_reg_reg g) env); intros. +case_eq (regmreg_edge_type_partition (interf_reg_mreg g) env); intros. +case_eq (regmreg_edge_type_partition (pref_reg_mreg g) env); intros. +simpl. +replace p2 with (fst (regmreg_edge_type_partition (interf_reg_mreg g) env)). +apply regmreg_in_snd_partition. auto. +assumption. +assumption. +rewrite H7. auto. +apply (proj2 (In_graph_edge_in_ext (Regs.P (fst x), Regs.M (snd x), None) + _ (or_intror _ H1))). +assumption. + +subst. elim H8. apply Regs.eq_refl. reflexivity. + +assert (OptionReg.eq (IRC float_graph float_palette (Regs.M (snd x))) (Some (Regs.M (snd x)))) as Hsnd. +apply mreg_refl_coloring. subst. +apply (proj2 (precolored_equiv _ _)). +unfold is_precolored. simpl. +split. auto. + +assert (EdgeSet.In (Regs.reg_to_Reg (fst x), Regs.mreg_to_Reg (snd x), None) + (IE float_graph)). +unfold int_graph. +apply regmreg_IE_translation. destruct x. simpl. +case_eq (regreg_edge_type_partition (interf_reg_reg g) env); intros. +case_eq (regreg_edge_type_partition (pref_reg_reg g) env); intros. +case_eq (regmreg_edge_type_partition (interf_reg_mreg g) env); intros. +case_eq (regmreg_edge_type_partition (pref_reg_mreg g) env); intros. +simpl. +replace p2 with (fst (regmreg_edge_type_partition (interf_reg_mreg g) env)). +apply regmreg_in_snd_partition. auto. +assumption. +assumption. +rewrite H4. auto. +apply (proj2 (In_graph_edge_in_ext (Regs.P (fst x), Regs.M (snd x), None) + _ (or_intror _ H1))). +assumption. + +rewrite H2 in Hsnd. inversion Hsnd. + +destruct (Loc.eq (S (Local (Zpos (fst x)) Tfloat)) (R (snd x))). +inversion e. reflexivity. +Qed. + +Import Registers. + +Lemma in_palette_not_in_temporaries : forall x, +VertexSet.In (Regs.M x) int_palette -> +~In (R x) temporaries. + +Proof. +unfold int_palette. intros x H. +repeat (destruct (proj1 (Props.Dec.F.add_iff _ _ _) H);[ + inversion H0; subst; intro H1; inversion H1; + [inversion H2| repeat (destruct H2 as [H2|H2];[inversion H2|])]; + assumption + | generalize H0; clear H H0; intro H]). +elim (VertexSet.empty_1 H). +Qed. + +Lemma in_palette_not_in_temporaries_float : forall x, +VertexSet.In (Regs.M x) float_palette -> +~In (R x) temporaries. + +Proof. +unfold int_palette. intros x H. +repeat (destruct (proj1 (Props.Dec.F.add_iff _ _ _) H);[ + inversion H0; subst; intro H1; inversion H1; + [inversion H2| repeat (destruct H2 as [H2|H2];[inversion H2|])]; + assumption + | generalize H0; clear H H0; intro H]). +elim (VertexSet.empty_1 H). +Qed. + +Lemma coloring_acceptable_loc : forall x, +loc_is_acceptable (my_graph_coloring g env x) = true. + +Proof. +intro x. +unfold loc_is_acceptable. +unfold my_graph_coloring in *. +change (snd (fst (fst (Typed_interfgraphs g env)))) with intG. +change (fst (fst (fst (Typed_interfgraphs g env)))) with intR. +change (snd (fst (Typed_interfgraphs g env))) with floatR. +change (snd (Typed_interfgraphs g env)) with floatG. +fold int_graph; fold float_graph. +unfold graph_coloring_aux. +change (map_to_coloring (IRC_map (graph_to_IRC_graph int_graph int_palette))) with + (IRC int_graph int_palette). +change (map_to_coloring (IRC_map (graph_to_IRC_graph float_graph float_palette))) with + (IRC float_graph float_palette). +case_eq (Regs.get_type (Regs.P x) env); intros. +case_eq (IRC int_graph int_palette (Regs.P x)); intros. +unfold IRC in *. +generalize (register_heuristic_mreg _ _ H0). intro H2. +destruct H2. rewrite H1. +destruct (List.In_dec Loc.eq (R x0) temporaries). +assert (~In (R x0) temporaries). +apply in_palette_not_in_temporaries. +rewrite <-H1. +generalize (proper_coloring_IRC_aux int_graph int_palette + (correct_palette_int)). intro H3. +unfold proper_coloring in H3. +do 2 destruct H3 as [_ H3]. +unfold proper_coloring_3 in H3. +apply (H3 (Regs.P x) t0). +unfold IRC in *. rewrite H0. apply OptionReg.eq_refl. +elim (H2 i). +reflexivity. +unfold IRC in *. auto. +case_eq (IRC float_graph float_palette (Regs.P x)); intros. +unfold IRC in *. +generalize (register_heuristic_mreg_float _ _ H0). intro H2. +destruct H2. rewrite H1. +destruct (List.In_dec Loc.eq (R x0) temporaries). +assert (~In (R x0) temporaries). +apply in_palette_not_in_temporaries_float. +rewrite <-H1. +generalize (proper_coloring_IRC_aux float_graph float_palette + correct_palette_float). intro H3. +unfold proper_coloring in H3. +do 2 destruct H3 as [_ H3]. +unfold proper_coloring_3 in H3. +apply (H3 (Regs.P x) t0). +unfold IRC in *. rewrite H0 in *. apply OptionReg.eq_refl. +elim (H2 i). +reflexivity. +unfold IRC in *. auto. +Qed. + +Lemma correct_alloc_3 : check_coloring_3 (all_interf_regs g) env (my_graph_coloring g env) = true. + +Proof. +unfold check_coloring_3. +apply Regset.for_all_1. + +unfold compat_bool. +intros. subst. reflexivity. + +unfold Regset.For_all. +intros x H. +rewrite coloring_acceptable_loc. simpl. +unfold same_typ. +rewrite <-Loc_reg_eq_type. +simpl. destruct (env x); reflexivity. + +Qed. + +Theorem correct_alloc : check_coloring g env (all_interf_regs g) (my_graph_coloring g env) = true. + +Proof. +unfold check_coloring. +rewrite correct_alloc_1. +rewrite correct_alloc_2. +rewrite correct_alloc_3. +auto. +Qed. + +End Coloring_to_allocation. + +Lemma precolored_sub_int_palette : forall x g env, +VertexSet.In x (precolored (int_graph g env)) -> VertexSet.In x int_palette. + +Proof. +Admitted. + +Lemma precolored_sub_float_palette : forall x g env, +VertexSet.In x (precolored (float_graph g env)) -> VertexSet.In x float_palette. + +Proof. +Admitted. + +Theorem allocation_correct : forall g env, +check_coloring g env (all_interf_regs g) (my_graph_coloring g env) = true. + +Proof. +intros. apply correct_alloc. +intros. apply (precolored_sub_int_palette x g env). assumption. +intros. apply (precolored_sub_float_palette x g env). assumption. +Qed. diff --git a/backend/MyRegisters.v b/backend/MyRegisters.v new file mode 100755 index 00000000..c2c7460d --- /dev/null +++ b/backend/MyRegisters.v @@ -0,0 +1,143 @@ +Require Import OrderedType. +Require Import Registers. +Require Import NArith. +Require Import Locations. +Require Import InterfGraph. + +Definition op_plus (x y : option N) : option N := +match (x,y) with +|(None,None) => None +|(Some a,Some b) => Some (Nplus a b) +|(Some a,None) => Some a +|(None,Some b) => Some b +end. + +(* Definition of registers as Pseudo-registers or machine registers *) + +Module Import Regs <: OrderedType. + +Inductive registers : Type := +|P : reg -> registers +|M : mreg -> registers. + +Definition t := registers. + +Definition reg_to_Reg := fun r => P r. +Definition mreg_to_Reg := fun r => M r. + +Inductive req : t -> t -> Prop := +|Peq : forall x y, eq x y -> req (P x) (P y) +|Meq : forall x y, eq x y -> req (M x) (M y). + +Definition eq := req. + +Inductive rlt : t -> t -> Prop := +|Plt : forall x y, OrderedReg.lt x y -> rlt (P x) (P y) +|Mlt : forall x y, OrderedMreg.lt x y -> rlt (M x) (M y) +|PMlt : forall x y, rlt (M x) (P y). + +Definition lt := rlt. + +Lemma eq_refl : forall x, eq x x. + +Proof. +induction x;constructor;reflexivity. +Qed. + +Lemma eq_sym : forall x y, eq x y -> eq y x. + +Proof. +induction 1;rewrite H;apply eq_refl. +Qed. + +Lemma eq_trans : forall x y z, eq x y -> eq y z -> eq x z. + +Proof. +intros x y z H H0. +inversion H;inversion H0;subst. +rewrite H5;apply eq_refl. +inversion H5. +inversion H5. +rewrite H5;apply eq_refl. +Qed. + +Lemma lt_trans : forall x y z, lt x y -> lt y z -> lt x z. + +Proof. +intros x y z H H0. +inversion H;inversion H0;subst. +inversion H5;subst. +constructor. apply (OrderedReg.lt_trans _ _ _ H1 H4). +inversion H5. +inversion H4. +inversion H5. +inversion H5;subst. +constructor. apply (OrderedMreg.lt_trans _ _ _ H1 H4). +constructor. +constructor. +inversion H4. +constructor. +Qed. + +Lemma lt_not_eq : forall x y, lt x y -> ~eq x y. + +Proof. +induction 1;intro H0;inversion H0. +elim (OrderedReg.lt_not_eq _ _ H H3). +elim (OrderedMreg.lt_not_eq _ _ H H3). +Qed. + +Lemma compare : forall x y, Compare lt eq x y. + +Proof. +intros x y. +destruct x;destruct y. +destruct (OrderedReg.compare r r0). +apply LT. constructor. assumption. +apply EQ. constructor. assumption. +apply GT. constructor. assumption. +apply GT. constructor. +apply LT. constructor. +destruct (OrderedMreg.compare m m0). +apply LT. constructor. assumption. +apply EQ. constructor. assumption. +apply GT. constructor. assumption. +Qed. + +Lemma eq_dec : forall x y, {eq x y}+{~eq x y}. + +Proof. +intros x y. +destruct (compare x y). +right. apply lt_not_eq. assumption. +left. assumption. +right. (cut (~ eq y x)). +intros H H0. elim H. apply eq_sym. assumption. +apply lt_not_eq. assumption. +Qed. + +Definition get_type x env := +match x with +| P y => env y +| M y => mreg_type y +end. + +End Regs. + +Definition is_mreg x := +match x with +| P _ => false +| M _ => true +end. + +Lemma mreg_ext : forall x y, +Regs.eq x y -> is_mreg x = is_mreg y. + +Proof. +intros. +destruct x; destruct y; simpl in *. +reflexivity. +inversion H. +inversion H. +reflexivity. +Qed. diff --git a/backend/Order_arith.v b/backend/Order_arith.v new file mode 100755 index 00000000..298468c2 --- /dev/null +++ b/backend/Order_arith.v @@ -0,0 +1,349 @@ +Require Import Arith. +Require Import ZArith. + +(* Definition of some facts other arith, for the termination order... + many uninteresting but necessary lemmas *) + +Fixpoint puissance_aux x y res {struct y} : nat := +match y with +| 0 => res +| S n => puissance_aux x n (x*res) +end. + +Definition puissance x y := puissance_aux x y 1. + +Lemma dec_base_aux : forall a b c d x, +0 < a -> +a + b +c +d < x -> +a*puissance x 3 + b*puissance x 2 + +c*x + d > 0. + +Proof. +intros a b c d x Ha H. +induction a. +inversion Ha. +unfold puissance;simpl;destruct x. +inversion H. +apply lt_le_trans with (m:=S x * (S x * (S x *1))). +rewrite mult_1_r. +case_eq (S x * (S x * S x)). +intro H0. +inversion H0. +intros n H0. +intuition. +intuition. +Qed. + +Lemma inf_diff : forall x y, +x - y > 0 -> x > y. + +Proof. +intuition. +Qed. + +Lemma plop : forall x y, +y > 0 -> +x <= x*y. + +Proof. +induction y;intros. +inversion H. +rewrite mult_succ_r. +intuition. +Qed. + +Lemma mult_zero : forall x y, +x*y = 0 -> x = 0 \/ y = 0. + +Proof. +intros. +induction x. +left;reflexivity. +induction y. +right;reflexivity. +inversion H. +Qed. + +Lemma mult_plop : forall x y z, +x > 0 -> +y < z -> +y*x < z*x. + +Proof. +intros x y z H H0. +induction x. +inversion H. +destruct (lt_eq_lt_dec x 0). +destruct s. +inversion l. +do 2 rewrite mult_succ_r. +rewrite e. +do 2 rewrite mult_0_r. +intuition. +generalize (IHx l). +do 2 rewrite mult_succ_r. +intuition. +Qed. + +Lemma mult_plop_eq : forall x y z, +x > 0 -> +y <= z -> +y*x <= z*x. + +Proof. +intros x y z H H0. +induction x. +inversion H. +destruct (lt_eq_lt_dec x 0). +destruct s. +inversion l. +do 2 rewrite mult_succ_r. +rewrite e. +do 2 rewrite mult_0_r. +intuition. +generalize (IHx l). +do 2 rewrite mult_succ_r. +intuition. +Qed. + +Lemma mult_plop_eq2 : forall x y z, +x > 0 -> +y <= z -> +x*y <= x*z. + +Proof. +intros x y z H H0. +induction x. +inversion H. +destruct (lt_eq_lt_dec x 0). +destruct s. +inversion l. +do 2 rewrite mult_succ_l. +rewrite e. +do 2 rewrite mult_0_l. +intuition. +generalize (IHx l). +do 2 rewrite mult_succ_l. +intuition. +Qed. + +Lemma add_inf : forall x y z, +x <= y -> z+x <= z+y. + +Proof. +intuition. +Qed. + +Lemma add_hd_eq : forall x y z, +x = y -> z+x = z+y. + +Proof. +auto. +Qed. + +Lemma add_hd_inf : forall x y z, +x < y -> z+x < z + y. + +Proof. +intuition. +Qed. + +Lemma dec_base : forall a b c d a' b' c' d' x, +a + b + c + d < x -> x > 0 -> +(a < a' -> +(a *puissance x 3 +b *puissance x 2 + c *x + d < + a' *puissance x 3 +b' *puissance x 2 + c' *x + d')). + +Proof. +intros a b c d a' b' c' d' x H H0 H1. +unfold puissance;simpl. +rewrite mult_1_r. +apply lt_le_trans with (m := (S a)*x*x*x). +rewrite mult_succ_l. +do 2 rewrite mult_plus_distr_r. +replace (a*(x*(x*x))+b*(x*x)+c*x+d) with (a*x*x*x+(b*x*x+c*x+d)). +apply add_hd_inf. +assert (c*x <= c*x*x). +apply plop. +assumption. +assert (d <= d*x*x). +apply le_trans with (m := d*x). +apply plop. +assumption. +apply plop. +assumption. +apply le_lt_trans with (b*x*x+c*x*x+d*x*x). +intuition. +replace (b*x*x+c*x*x+d*x*x) with ((b+c+d)*x*x). +apply mult_plop. +assumption. +apply mult_plop. +assumption. +intuition. +do 4 rewrite mult_plus_distr_r. +reflexivity. +do 2 rewrite plus_assoc. +replace (a*(x* (x*x))) with (a*x*x*x). +replace (b*(x*x)) with (b*x*x). +reflexivity. +intuition. +replace (x*(x*x)) with (x*x*x). +rewrite mult_assoc. +intuition. +intuition. +apply le_trans with (m := a'*(x*(x*x))). +do 2 rewrite mult_assoc. +intuition. +apply mult_plop_eq. +assumption. +apply mult_plop_eq. +assumption. +apply mult_plop_eq. +assumption. +assumption. +intuition. +Qed. + +Lemma dec_base2 : forall a b c a' b' c' x, +a + b + c < x -> x > 0 -> +(a < a' -> +(a *puissance x 2 +b *x+ c < + a' *puissance x 2 +b' *x + c')). + +Proof. +intros a b c a' b' c' x H H0 H1. +unfold puissance;simpl. +rewrite mult_1_r. +apply lt_le_trans with (m := S a*x*x). +rewrite mult_succ_l. +rewrite mult_plus_distr_r. +rewrite mult_assoc. +rewrite <-plus_assoc. +apply add_hd_inf. +assert (c <= c*x). +apply plop. +assumption. +apply le_lt_trans with (m := b*x+c*x). +intuition. +replace (b*x+c*x) with ((b+c)*x). +apply mult_plop. +assumption. +intuition. +intuition. +apply le_trans with (m := a'*x*x). +apply mult_plop_eq. +assumption. +apply mult_plop_eq. +assumption. +intuition. +rewrite mult_assoc. +intuition. +Qed. + +Lemma dec_base3 : forall a b a' b' x, +a + b < x -> x > 0 -> a < a' +-> (a *x +b < a' *x + b'). + +Proof. +intros a b a' b' x H H0 H1. +apply lt_le_trans with (m := S a*x). +rewrite mult_succ_l. +intuition. +assert (S a <= a'). +intuition. +apply le_trans with (m:= a'*x). +apply mult_plop_eq;assumption. +intuition. +Qed. + +Lemma puiss_aux : forall x y n, +puissance_aux x y n = n*puissance x y. + +Proof. +unfold puissance;induction y;simpl. +intuition. +intro n. +rewrite mult_1_r. +rewrite (IHy x). +rewrite (IHy (x*n)). +rewrite mult_assoc. +rewrite (mult_comm x n). +reflexivity. +Qed. + +Lemma puiss_step : forall x n, +puissance x (S n) = x*puissance x n. + +Proof. +intros. +unfold puissance;simpl. +rewrite puiss_aux;simpl. +intuition. +Qed. + +Lemma pos_puiss : forall x y, +x > 0 -> puissance x y > 0. + +Proof. +induction y;intros. +unfold puissance;simpl. +intuition. +rewrite puiss_step. +generalize (IHy H);intro H0. +case_eq (x*puissance x y);intros. +generalize (mult_zero _ _ H1);intros. +destruct H2;intuition. +intuition. +Qed. + +Lemma incr_puis : forall x y n, +x <= y -> +puissance x n <= puissance y n. + +Proof. +induction n;intros;simpl. +intuition. +do 2 rewrite puiss_step. +destruct x. +intuition. +destruct y. +inversion H. +apply le_trans with (m := S x*puissance (S y) n). +apply mult_plop_eq2. +intuition. +apply (IHn H). +apply mult_plop_eq. +apply pos_puiss. +intuition. +assumption. +Qed. + +Lemma incr_puis2 : forall x n p, +n <= p -> +x > 0 -> +puissance x n <= puissance x p. + +Proof. +induction n;intros. +unfold puissance;simpl. +fold (puissance x p). +generalize (pos_puiss). +intuition. +unfold puissance;simpl. +rewrite mult_1_r. +do 2 rewrite puiss_aux. +rewrite mult_1_l. +generalize (IHn (p-1));intro. +assert (puissance x n <= puissance x (p-1)). +apply H1. +intuition. +assumption. +destruct p. +inversion H. +rewrite puiss_step. +apply mult_plop_eq2. +assumption. +apply IHn. +intuition. +assumption. +Qed. \ No newline at end of file diff --git a/backend/OrderedOption.v b/backend/OrderedOption.v new file mode 100755 index 00000000..b2cebe2b --- /dev/null +++ b/backend/OrderedOption.v @@ -0,0 +1,96 @@ +Require Import FSets. + +(* Definition of options types as ordered types, + in order to define weights *) + +Module OrderedOpt (O : OrderedType) <: OrderedType. + +Definition t := option O.t. + +Inductive eq_ : t -> t -> Prop := +|None_eq : eq_ None None +|Some_eq : forall x y, O.eq x y -> eq_ (Some x) (Some y). + +Definition eq := eq_. + +Inductive lt_ : t -> t -> Prop := +|None_lt : forall o : O.t, lt_ None (Some o) +|Some_lt : forall o o', O.lt o o' -> lt_ (Some o) (Some o'). + +Definition lt := lt_. + +Lemma eq_refl : forall x : t, eq x x. + +Proof. +unfold eq;intro x;destruct x;constructor;apply O.eq_refl. +Qed. + +Lemma eq_sym : forall x y : t, eq x y -> eq y x. + +Proof. +unfold eq;intros x y H;destruct x;destruct y. +constructor;inversion H;apply O.eq_sym;assumption. +inversion H. +inversion H. +assumption. +Qed. + +Lemma eq_trans : forall x y z, eq x y -> eq y z -> eq x z. + +Proof. +unfold eq;intros x y z H H0. +inversion H;inversion H0. +constructor. +rewrite <-H2 in H4;inversion H4. +rewrite <-H3 in H4;inversion H4. +rewrite <-H3 in H5;inversion H5;subst. +constructor;apply (O.eq_trans H1 H4). +Qed. + +Lemma lt_trans : forall x y z, lt x y -> lt y z -> lt x z. + +Proof. +unfold lt;intros x y z H H0. +inversion H;inversion H0;constructor. +rewrite <-H3 in H4;inversion H4. +rewrite <-H3 in H5;inversion H5. +subst. +apply (O.lt_trans H1 H4). +Qed. + +Lemma lt_not_eq : forall x y, lt x y -> ~eq x y. + +Proof. +intros x y H. +unfold eq;intro Heq. +inversion H;inversion Heq. +rewrite <-H3 in H1;inversion H1. +rewrite <-H0 in H3;inversion H3. +rewrite <-H1 in H3;inversion H3. +subst;inversion H4;inversion H5;subst. +elim (O.lt_not_eq H0 H3). +Qed. + +Lemma compare : forall x y : t, Compare lt eq x y. + +Proof. +intros x y. +destruct x;destruct y. +destruct (O.compare t0 t1); +[apply LT;unfold lt|apply EQ;unfold eq|apply GT;unfold lt];constructor;assumption. +apply GT;unfold lt;constructor. +apply LT;unfold lt;constructor. +apply EQ;unfold eq;constructor. +Qed. + +Lemma eq_dec : forall x y, {eq x y}+{~eq x y}. + +Proof. +intros x y. +destruct (compare x y). +right. apply lt_not_eq. assumption. +left. assumption. +right. intro H. generalize (eq_sym _ _ H). intro H0. elim (lt_not_eq _ _ l H0). +Qed. + +End OrderedOpt. \ No newline at end of file diff --git a/backend/Regs.v b/backend/Regs.v new file mode 100755 index 00000000..33e3fea9 --- /dev/null +++ b/backend/Regs.v @@ -0,0 +1,122 @@ +Require Import Registers. + + +Definition op_plus (x y : option N) : option N := +match (x,y) with +|(None,None) => None +|(Some a,Some b) => Some (Nplus a b) +|(Some a,None) => Some a +|(None,Some b) => Some b +end. + +Module Regs <: OrderedType. + +(* Definition of registers as Pseudo-registers or machine registers *) + +Inductive registers : Type := +|P : reg -> registers +|M : mreg -> registers. + +Definition t := registers. + +Definition reg_to_Reg := fun r => P r. +Definition mreg_to_Reg := fun r => M r. + +Inductive req : t -> t -> Prop := +|Peq : forall x y, eq x y -> req (P x) (P y) +|Meq : forall x y, eq x y -> req (M x) (M y). + +Definition eq := req. + +Inductive rlt : t -> t -> Prop := +|Plt : forall x y, OrderedReg.lt x y -> rlt (P x) (P y) +|Mlt : forall x y, OrderedMreg.lt x y -> rlt (M x) (M y) +|PMlt : forall x y, rlt (M x) (P y). + +Definition lt := rlt. + +Lemma eq_refl : forall x, eq x x. + +Proof. +induction x;constructor;reflexivity. +Qed. + +Lemma eq_sym : forall x y, eq x y -> eq y x. + +Proof. +induction 1;rewrite H;apply eq_refl. +Qed. + +Lemma eq_trans : forall x y z, eq x y -> eq y z -> eq x z. + +Proof. +intros x y z H H0. +inversion H;inversion H0;subst. +rewrite H5;apply eq_refl. +inversion H5. +inversion H5. +rewrite H5;apply eq_refl. +Qed. + +Lemma lt_trans : forall x y z, lt x y -> lt y z -> lt x z. + +Proof. +intros x y z H H0. +inversion H;inversion H0;subst. +inversion H5;subst. +constructor. apply (OrderedReg.lt_trans H1 H4). +inversion H5. +inversion H4. +inversion H5. +inversion H5;subst. +constructor. apply (OrderedMreg.lt_trans _ _ _ H1 H4). +constructor. +constructor. +inversion H4. +constructor. +Qed. + +Lemma lt_not_eq : forall x y, lt x y -> ~eq x y. + +Proof. +induction 1;intro H0;inversion H0. +elim (OrderedReg.lt_not_eq H H3). +elim (OrderedMreg.lt_not_eq _ _ H H3). +Qed. + +Lemma compare : forall x y, Compare lt eq x y. + +Proof. +intros x y. +destruct x;destruct y. +destruct (OrderedReg.compare r r0). +apply LT. constructor. assumption. +apply EQ. constructor. assumption. +apply GT. constructor. assumption. +apply GT. constructor. +apply LT. constructor. +destruct (OrderedMreg.compare m m0). +apply LT. constructor. assumption. +apply EQ. constructor. assumption. +apply GT. constructor. assumption. +Qed. + +Lemma eq_dec : forall x y, {eq x y}+{~eq x y}. + +Proof. +intros x y. +destruct (compare x y). +right. apply lt_not_eq. assumption. +left. assumption. +right. (cut (~ eq y x)). +intros H H0. elim H. apply eq_sym. assumption. +apply lt_not_eq. assumption. +Qed. + +Definition get_type x env := +match x with +| P y => env y +| M y => mreg_type y +end. + +End Regs. \ No newline at end of file diff --git a/backend/Remove_Vertex_Adjacency.v b/backend/Remove_Vertex_Adjacency.v new file mode 100755 index 00000000..668063d1 --- /dev/null +++ b/backend/Remove_Vertex_Adjacency.v @@ -0,0 +1,73 @@ +Require Import FSets. +Require Import InterfGraphMapImp. +Require Import Interference_adjacency. +Require Import Edges. +Require Import MyRegisters. + +Import RegFacts Props. + +Module Register := Regs. + +(* For any vertex x different from r, the neighborhood of x in + (remove_vertex r g) is its one in g, minus r *) +Lemma remove_interf_adj : forall x r g, +~Register.eq x r -> +VertexSet.Equal (interference_adj x (remove_vertex r g)) + (VertexSet.remove r (interference_adj x g)). + +Proof. +intros. +split; intros. +apply VertexSet.remove_2. intro. +rewrite <-H1 in H0. rewrite in_interf in H0. +generalize (proj1 (In_graph_edge_in_ext _ _ H0)). change_rewrite. intro. +rewrite In_remove_vertex in H2. destruct H2. elim (H3 (Register.eq_refl _)). +rewrite in_interf in H0. rewrite In_remove_edge in H0. destruct H0. +rewrite in_interf. assumption. + +rewrite in_interf. rewrite In_remove_edge. +split. +rewrite <-in_interf. apply (VertexSet.remove_3 H0). +intro. destruct H1; change_rewrite. +elim (VertexSet.remove_1 H1 H0). +elim H. auto. +Qed. + +(* If x is different from r and does not belong to the interference neighborhood + of r in g, then the interference neighborhood of x in (remove_vertex r g) + is equal to the interference neighborhood of x in g *) +Lemma interf_adj_remove : forall x r g, +~VertexSet.In x (interference_adj r g) -> +~Register.eq x r -> +VertexSet.Equal (interference_adj x g) (interference_adj x (remove_vertex r g)). + +Proof. +intros x r g H H0. +rewrite remove_interf_adj. rewrite remove_equal. apply VertexSet.eq_refl. +intro. elim H. apply interf_adj_comm. assumption. +auto. +Qed. + +(* The interference neighborhood of x in (remove_vertex r g) + is a subset of the interference neighborhood of x in g *) +Lemma sub_remove_interf : forall x r g, +~Register.eq x r -> +VertexSet.Subset (interference_adj x (remove_vertex r g)) + (interference_adj x g). + +Proof. +intros x r g H. rewrite remove_interf_adj. +unfold VertexSet.Subset. intros y H0. +apply (VertexSet.remove_3 H0). +assumption. +Qed. + +(* If x is a neighbor of r in g, then x belongs to (remove_vertex r g) *) +Lemma in_interf_adj_in_remove : forall x r g, +VertexSet.In x (interference_adj r g) -> In_graph x (remove_vertex r g). + +Proof. +intros x r g H. rewrite In_remove_vertex. split. +rewrite in_interf in H. apply (proj1 (In_graph_edge_in_ext _ _ H)). +intro H0. rewrite H0 in H. elim (not_in_interf_self _ _ H). +Qed. diff --git a/backend/Remove_Vertex_Degree.v b/backend/Remove_Vertex_Degree.v new file mode 100755 index 00000000..e296265c --- /dev/null +++ b/backend/Remove_Vertex_Degree.v @@ -0,0 +1,173 @@ +Require Import FSets. +Require Import InterfGraphMapImp. +Require Import Remove_Vertex_Adjacency. +Require Import ZArith. +Require Import Edges. +Require Import MyRegisters. +Require Import Interference_adjacency. + +Module Register := Regs. + +Import RegFacts Props. + +(* The degree of any vertex different from r decreases when + r is removed from the a graph g. Hence, a low-degree vertex of g + is a low-degree vertex of (remove_vertex r g) *) +Lemma low_remove_low : forall x r g K, +has_low_degree g K x = true -> +~Register.eq x r -> +has_low_degree (remove_vertex r g) K x = true. + +Proof. +intros x r g K H H0. unfold has_low_degree, interf_degree in *. +generalize (sub_remove_interf _ _ g H0);intro H1. +generalize (subset_cardinal H1);intro H2. +destruct (le_lt_dec K (VertexSet.cardinal (interference_adj x g))). +inversion H. +destruct (le_lt_dec K (VertexSet.cardinal (interference_adj x (remove_vertex r g)))). +apply False_ind;intuition. +auto. +Qed. + +(* For the same reason, a high-degree vertex of (remove_vertex r g) + is a high-degree vertex of g *) +Lemma not_low_remove_not_low : forall x r g K, +has_low_degree (remove_vertex r g) K x = false -> +~Register.eq x r -> +has_low_degree g K x = false. + +Proof. +intros x r g K H H0. +case_eq (has_low_degree g K x);[intros|auto]. +rewrite (low_remove_low _ _ _ _ H1) in H. inversion H. +auto. +Qed. + +(* The degree of any vertex x which is not an interference neighbor of + (remove_vertex r g) is the same in g and in (remove_vertex r g) *) +Lemma low_deg : forall x r g K, +~VertexSet.In x (interference_adj r g) -> +~Register.eq x r -> +has_low_degree g K x = has_low_degree (remove_vertex r g) K x. + +Proof. +intros x r g K H H0. unfold has_low_degree, interf_degree. +destruct (le_lt_dec K (VertexSet.cardinal (interference_adj x g))); +destruct (le_lt_dec K (VertexSet.cardinal (interference_adj x (remove_vertex r g)))). +reflexivity. +rewrite <-(cardinal_m(interf_adj_remove _ _ _ H H0)) in l0. +apply False_ind. intuition. +elim (lt_irrefl (VertexSet.cardinal (interference_adj x g))). +apply lt_le_trans with (m := K). auto. +rewrite (cardinal_m (interf_adj_remove x r g H H0)). auto. +reflexivity. +Qed. + +Lemma lt_le_S_eq : forall n x, +x < n -> +n <= S x -> +n = S x. + +Proof. +induction n; intros. +intuition. +induction x; intros. +apply eq_S. intuition. +apply eq_S. apply IHn. intuition. intuition. +Qed. + +Lemma le_S_irrefl : forall x, +S x <= x -> False. + +Proof. +induction x; intros. +inversion H. +apply IHx. apply le_S_n. assumption. +Qed. + +(* If x is a low-degree degree of (remove_vertex r g) and x is + a high-degree vertex of g then the interference degree + of x in g is exactly k *) +Lemma degree_dec_remove_K : forall x k r g, +~Register.eq x r -> +has_low_degree g k x = false -> +has_low_degree (remove_vertex r g) k x = true -> +interf_degree g x = k. + +Proof. +unfold has_low_degree, interf_degree. intros. +destruct (le_lt_dec k (VertexSet.cardinal (interference_adj x (remove_vertex r g)))) in H1. +inversion H1. +rewrite (cardinal_m (remove_interf_adj _ _ _ H)) in l. +unfold has_low_degree in H0. +destruct (le_lt_dec k (VertexSet.cardinal (interference_adj x g))). +destruct (In_dec r (interference_adj x g)). +generalize (remove_cardinal_1 i). intro HH. rewrite <-HH in l0. +set (card := VertexSet.cardinal (VertexSet.remove r (interference_adj x g))) in *. +rewrite <-HH. symmetry. apply lt_le_S_eq; auto. +generalize (remove_cardinal_2 n). intro HH. rewrite <-HH in l0. +elim (lt_irrefl k). intuition. +inversion H0. +Qed. + +(* Any x which is of high-degree in g and of low-degree in (remove_vertex r g) + interferes with r *) +Lemma low_degree_in_interf : forall x r g K, +has_low_degree (remove_vertex r g) K x = true -> +~Register.eq x r -> +has_low_degree g K x = false -> +VertexSet.In x (interference_adj r g). + +Proof. +unfold has_low_degree, interf_degree. intros x r g K H HH H0. +destruct (le_lt_dec K (VertexSet.cardinal (interference_adj x g))). +destruct (le_lt_dec K (VertexSet.cardinal (interference_adj x (remove_vertex r g)))). +inversion H. +destruct (In_dec x (interference_adj r g)). +assumption. +generalize (interf_adj_remove _ _ _ n HH);intro H2. +rewrite (cardinal_m H2) in l. +apply False_ind. intuition. +inversion H0. +Qed. + +(* Reciprocally, a high-degree vertex x of g which is + exactly of degree k in g and interferes with r is + of low-degree in (remove_vertex r g) *) + +Lemma degree_K_remove_dec : forall x k r g, +~Register.eq x r -> +interf_degree g x = k -> +VertexSet.In x (interference_adj r g) -> +has_low_degree (remove_vertex r g) k x = true. + +Proof. +unfold has_low_degree, interf_degree. intros. +destruct (le_lt_dec k (VertexSet.cardinal (interference_adj x g))). +destruct (le_lt_dec k (VertexSet.cardinal (interference_adj x (remove_vertex r g)))). +rewrite (cardinal_m (remove_interf_adj _ _ _ H)) in l0. +generalize (remove_cardinal_1 (interf_adj_comm _ _ _ H1)). +intro HH. rewrite <-HH in l. +set (card := VertexSet.cardinal (VertexSet.remove r (interference_adj x g))) in *. +rewrite <-HH in H0. rewrite <-H0 in *. elim (le_S_irrefl _ l0). +reflexivity. +apply False_ind. intuition. +Qed. + +(* Finally, an unused but meaningful theorem summarizing + conditions leading to an evolution of the interference degrees + when a vertex r is removed *) +Theorem remove_low_degree_evolution : forall x k r g, +~Register.eq x r -> +((has_low_degree g k x = false /\ has_low_degree (remove_vertex r g) k x = true) + <-> + (VertexSet.In x (interference_adj r g) /\ interf_degree g x = k)). + +Proof. +intros. split; intros; destruct H0. +split. eapply low_degree_in_interf; eauto. eapply degree_dec_remove_K; eauto. +cut (has_low_degree g k x = false). intro. +split. assumption. +apply degree_K_remove_dec; auto. +unfold has_low_degree. rewrite H1. destruct (le_lt_dec k k). auto. elim (lt_irrefl _ l). +Qed. diff --git a/backend/Remove_Vertex_Move.v b/backend/Remove_Vertex_Move.v new file mode 100755 index 00000000..a4734731 --- /dev/null +++ b/backend/Remove_Vertex_Move.v @@ -0,0 +1,206 @@ +Require Import FSets. +Require Import InterfGraphMapImp. +Require Import WS. +Require Import Edges. +Require Import MyRegisters. +Require Import Affinity_relation. +Require Import Interference_adjacency. + +Module Register := Regs. + +Import RegFacts Props Edge. + +(* A nonmove-related vertex of g is not move-related + after the removal of a vertex *) +Lemma move_remove : forall x r g, +move_related g x = false -> +move_related (remove_vertex r g) x = false. + +Proof. +intros. +apply move_related_false_charac2. intros. +apply move_related_false_charac with (g:=g); auto. +rewrite In_remove_edge in H1. intuition. +Qed. + +(* Equivalently, a move-related vertex of (remove_vertex r g) + is move-related in g *) +Lemma move_remove_2 : forall x r g, +move_related (remove_vertex r g) x = true -> +move_related g x = true. + +Proof. +intros x r g H. +generalize (move_related_charac _ _ H);intro H0. +destruct H0 as [y H0]. destruct H0 as [H0 H1]. destruct H1 as [H1 H2]. +apply move_related_charac2 with (e:= y). +assumption. +rewrite In_remove_edge in H1. destruct H1. assumption. assumption. +Qed. + +(* If x (different from r) is move-related in g and nonmove-related + in (remove_vertex r g) then x is a preference neighbor of r in g *) +Lemma move_related_not_remove_in_pref : forall x r g, +~Register.eq r x -> +move_related g x = true -> +move_related (remove_vertex r g) x = false -> +VertexSet.In x (preference_adj r g). + +Proof. +intros. +generalize (move_related_charac _ _ H0). intro. +destruct H2. destruct H2. destruct H3. +destruct (incident_dec x0 r). +destruct H5. destruct H4. +elim H. rewrite H4. rewrite H5. auto. +unfold aff_edge in H2. destruct H2. +rewrite in_pref. exists x1. +assert (eq (x,r,Some x1) x0). +rewrite edge_comm. apply eq_ordered_eq. +unfold E.eq; split; simpl. split; auto. +unfold get_weight in H2. rewrite H2. apply OptionN_as_OT.eq_refl. +rewrite H6. assumption. +destruct H4. +destruct H2. +rewrite in_pref. exists x1. +assert (eq (x,r,Some x1) x0). +apply eq_ordered_eq. +unfold E.eq; split; simpl. split; auto. +unfold get_weight in H2. rewrite H2. apply OptionN_as_OT.eq_refl. +rewrite H6. assumption. +elim H. rewrite H4. rewrite H5. auto. +assert (In_graph_edge x0 (remove_vertex r g)). +rewrite In_remove_edge. split; assumption. +generalize (Aff_edge_aff _ _ H6 H2). intro. +destruct H7. destruct H4. +rewrite (compat_bool_move _ _ _ (Register.eq_sym _ _ H4)) in H7. rewrite H1 in H7. inversion H7. +rewrite (compat_bool_move _ _ _ (Register.eq_sym _ _ H4)) in H8. rewrite H1 in H8. inversion H8. +Qed. + +(* The preference neighborhood of any vertex x different from r in + (remove_vertex r g) is obtained by removing r from the interference + neighborhood of x in g *) +Lemma remove_pref_adj : forall x r g, +~Register.eq x r -> +VertexSet.Equal (preference_adj x (remove_vertex r g)) + (VertexSet.remove r (preference_adj x g)). + +Proof. +intros. +split; intros. +apply VertexSet.remove_2. intro. +rewrite <-H1 in H0. rewrite in_pref in H0. destruct H0. +generalize (proj1 (In_graph_edge_in_ext _ _ H0)). change_rewrite. intro. +rewrite In_remove_vertex in H2. destruct H2. elim (H3 (Register.eq_refl _)). +rewrite in_pref. rewrite in_pref in H0. destruct H0. exists x0. +rewrite In_remove_edge in H0. destruct H0. assumption. + +generalize (VertexSet.remove_3 H0). intro. +rewrite in_pref in H1. destruct H1. +rewrite in_pref. exists x0. rewrite In_remove_edge. split. assumption. +intro. destruct H2; change_rewrite. +elim (VertexSet.remove_1 H2 H0). +elim H. intuition. +Qed. + +(* The preference degree of any vertex which is move-related + in g and nonmove-related in (remove_vertex r g) is equal to 1 *) +Lemma pref_degree_dec_remove_1 : forall x r g, +~Register.eq x r -> +move_related g x = true -> +move_related (remove_vertex r g) x = false -> +pref_degree g x = 1. + +Proof. +unfold pref_degree. intros. +generalize (not_move_related_empty_pref _ _ H1). intro. +generalize (remove_pref_adj x r g H). intro. +rewrite H2 in H3. +cut (~Register.eq r x). intro. +generalize (move_related_not_remove_in_pref _ _ _ H4 H0 H1). intro. +generalize (pref_adj_comm _ _ _ H5). intro. +generalize (remove_cardinal_1 H6). intro. +rewrite <-H7. apply eq_S. rewrite <-H3. +rewrite <-Props.cardinal_Empty. apply VertexSet.empty_1. +intuition. +Qed. + +Lemma cardinal_1_singleton : forall x s, +VertexSet.In x s -> +VertexSet.cardinal s = 1 -> +VertexSet.Equal s (VertexSet.singleton x). + +Proof. +intros. +apply VertexSet.eq_sym. +apply remove_singleton_empty. assumption. +assert (VertexSet.cardinal (VertexSet.remove x s) = 0). +rewrite <-remove_cardinal_1 with (x:=x) in H0. auto. +assumption. +rewrite <-Props.cardinal_Empty in H1. +rewrite (empty_is_empty_1 H1). apply VertexSet.eq_refl. +Qed. + +(* Reciprocally, any vertex different from r which has a preference + degree equal to 1 in g and is a preference neighbor of r in g is + nonmove-related in (remove_vertex r g) *) +Lemma pref_degree_1_remove_dec : forall x r g, +~Register.eq x r -> +pref_degree g x = 1 -> +VertexSet.In x (preference_adj r g) -> +move_related (remove_vertex r g) x = false. + +Proof. +intros. +case_eq (move_related (remove_vertex r g) x); intros. +generalize (remove_pref_adj x r g H). intro. +assert (VertexSet.Equal (preference_adj x g) (VertexSet.singleton r)). +apply cardinal_1_singleton. apply pref_adj_comm. assumption. assumption. +rewrite H4 in H3. +cut (VertexSet.Equal (preference_adj x (remove_vertex r g)) VertexSet.empty). intro. +generalize (move_related_charac _ _ H2). intro. +destruct H6. destruct H6. destruct H7. +destruct H8. +assert (VertexSet.In (snd_ext x0) (preference_adj x (remove_vertex r g))). +destruct H6. rewrite in_pref. exists x1. +assert (eq (snd_ext x0, x, Some x1) x0). +rewrite edge_comm. apply eq_ordered_eq. +unfold E.eq. split; intros. simpl. split; intuition. apply Regs.eq_refl. +auto. auto. simpl. unfold get_weight in H6. rewrite H6. apply OptionN_as_OT.eq_refl. +rewrite H9. assumption. +rewrite H5 in H9. elim (VertexSet.empty_1 H9). +assert (VertexSet.In (fst_ext x0) (preference_adj x (remove_vertex r g))). +destruct H6. rewrite in_pref. exists x1. +assert (eq (fst_ext x0, x, Some x1) x0). +apply eq_ordered_eq. +unfold E.eq. split; intros. simpl. split; intuition. apply Regs.eq_refl. +auto. auto. simpl. unfold get_weight in H6. rewrite H6. apply OptionN_as_OT.eq_refl. +rewrite H9. assumption. +rewrite H5 in H9. elim (VertexSet.empty_1 H9). +rewrite H3. +split; intros. +destruct (Register.eq_dec r a). +elim (VertexSet.remove_1 e H5). +generalize (VertexSet.remove_3 H5). intro. +generalize (VertexSet.singleton_1 H6). intro. +elim (n H7). +elim (VertexSet.empty_1 H5). +reflexivity. +Qed. + +(* Meaningful theorem *) +Theorem Remove_vertex_move_evolution : forall x r g, +~Register.eq x r -> +((move_related g x = true /\ move_related (remove_vertex r g) x = false) + <-> + (pref_degree g x = 1 /\ VertexSet.In x (preference_adj r g))). + +Proof. +split; intros. +destruct H0. +split. apply pref_degree_dec_remove_1 with (r:=r); auto. + apply move_related_not_remove_in_pref; auto. +destruct H0. +split. apply move_related_card. congruence. + apply pref_degree_1_remove_dec; auto. +Qed. diff --git a/backend/Remove_Vertex_WL.v b/backend/Remove_Vertex_WL.v new file mode 100755 index 00000000..2ad7c766 --- /dev/null +++ b/backend/Remove_Vertex_WL.v @@ -0,0 +1,478 @@ +Require Import FSets. +Require Import InterfGraphMapImp. +Require Import IRC_graph. +Require Import ZArith. +Require Import Edges. +Require Import Remove_Vertex_Degree. +Require Import WS. +Require Import Remove_Vertex_Move. +Require Import Affinity_relation. +Require Import Interference_adjacency. + +Import RegFacts Props OTFacts. + +Definition eq_K x K := match eq_nat_dec x K with +|left _ => true +|right _ => false +end. + +Lemma eq_K_1 : forall x y, +y = x -> +eq_K x y = true. + +Proof. +intros. unfold eq_K. rewrite <-H. destruct (eq_nat_dec y y); intuition. +Qed. + +Lemma eq_K_2 : forall x y, +eq_K x y = true -> +x = y. + +Proof. +intros. unfold eq_K in H. destruct (eq_nat_dec x y). auto. inversion H. +Qed. + +Lemma eq_K_compat : forall K g, +compat_bool Register.eq (fun x => eq_K K (VertexSet.cardinal (interference_adj x g))). + +Proof. +unfold compat_bool. intros. rewrite (compat_interference_adj _ _ _ H). reflexivity. +Qed. + +Lemma eq_K_compat_pref : forall K g, +compat_bool Register.eq (fun x => eq_K K (VertexSet.cardinal (preference_adj x g))). + +Proof. +unfold compat_bool. intros. rewrite (compat_preference_adj _ _ _ H). reflexivity. +Qed. + +Lemma compat_move_up : forall g K, +compat_bool Register.eq +(fun x => eq_K 1 (VertexSet.cardinal (preference_adj x g)) && has_low_degree g K x). + +Proof. +unfold compat_bool; intros. +rewrite (compat_preference_adj _ _ _ H). +destruct (eq_K 1 (VertexSet.cardinal (preference_adj y g))). simpl. +apply compat_bool_low. assumption. +simpl. reflexivity. +Qed. + +Definition remove_wl_2 r ircg K := + let g := irc_g ircg in + let wl := irc_wl ircg in + let simplify := get_simplifyWL wl in + let freeze := get_freezeWL wl in + let spillWL := get_spillWL wl in + let movesWL := get_movesWL wl in + let pre := precolored g in + let int_adj := interference_adj r g in + let not_pre_int_adj := VertexSet.diff int_adj pre in + let pre_adj := preference_adj r g in + let not_pre_pre_adj := VertexSet.diff pre_adj pre in + let newlow := VertexSet.filter (fun x => eq_K K (VertexSet.cardinal (interference_adj x g))) not_pre_int_adj in + let (free, simp) := VertexSet.partition (move_related g) newlow in + let newnmr := VertexSet.filter (fun x => eq_K 1 (VertexSet.cardinal (preference_adj x g)) && has_low_degree g K x) not_pre_pre_adj in + let simplifyWL__ := VertexSet.union simplify simp in + let simplifyWL_ := VertexSet.union simplifyWL__ newnmr in + let simplifyWL' := VertexSet.remove r simplifyWL_ in + let freezeWL__ := VertexSet.diff freeze newnmr in + let freezeWL_ := VertexSet.union freezeWL__ free in + let freezeWL' := VertexSet.remove r freezeWL_ in + let spillWL_ := VertexSet.diff spillWL newlow in + let spillWL' := VertexSet.remove r spillWL_ in + let movesWL' := not_incident_edges r movesWL g in + (spillWL', freezeWL', simplifyWL', movesWL'). + +Lemma WS_remove_wl_2 : forall r ircg, +WS_properties (remove_vertex r (irc_g ircg)) (VertexSet.cardinal (pal ircg)) + (remove_wl_2 r ircg (VertexSet.cardinal (pal ircg))). + +Proof. +unfold remove_wl_2. intros r ircg. +generalize (HWS_irc ircg). intro HWS. rewrite <-(Hk ircg) in *. +set (g' := remove_vertex r (irc_g ircg)) in *. +set (k := VertexSet.cardinal (pal ircg)) in *. +set (g := irc_g ircg) in *. +set (wl := irc_wl ircg) in *. +set ( simplify := get_simplifyWL wl ) in *. +set ( freeze := get_freezeWL wl ) in *. +set ( spillWL := get_spillWL wl ) in *. +set ( int_adj := interference_adj r g ) in *. +set ( not_pre_int_adj := VertexSet.diff int_adj (precolored g) ) in *. +set ( pre_adj := preference_adj r g ) in *. +set ( not_pre_pre_adj := VertexSet.diff pre_adj (precolored g) ) in *. +set ( low := VertexSet.filter (fun x => eq_K k (VertexSet.cardinal (interference_adj x g))) not_pre_int_adj ) in *. +set ( simpfree := VertexSet.partition (move_related g) low ) in *. +case_eq simpfree. intros free simp Hsf. +unfold simpfree in Hsf. +set ( nmr := VertexSet.filter (fun x => eq_K 1 (VertexSet.cardinal (preference_adj x g)) && has_low_degree g k x) not_pre_pre_adj) in *. +set ( simplifyWL__ := VertexSet.union simplify simp ) in *. +set ( simplifyWL_ := VertexSet.union simplifyWL__ nmr) in *. +set ( simplifyWL' := VertexSet.remove r simplifyWL_ ) in *. +set ( freezeWL__ := VertexSet.diff freeze nmr ) in *. +set ( freezeWL_ := VertexSet.union freezeWL__ free ) in *. +set ( freezeWL' := VertexSet.remove r freezeWL_ ) in *. +set ( spillWL_ := VertexSet.diff spillWL low ) in *. +set ( spillWL' := VertexSet.remove r spillWL_ ) in *. +set ( movesWL' := not_incident_edges r (get_movesWL wl) g) in *. + +unfold WS_properties. split. +split; intro H. +unfold get_spillWL in H. simpl in H. +(* spillWL' => *) +unfold spillWL' in H. +generalize (VertexSet.remove_3 H). intro H0. +unfold spillWL_ in H0. +generalize (VertexSet.diff_1 H0). intro H1. +generalize (VertexSet.diff_2 H0). intro H2. +split. +case_eq (has_low_degree g' k x); intros. +elim H2. apply VertexSet.filter_3. +apply eq_K_compat. +unfold not_pre_int_adj. apply VertexSet.diff_3. +apply low_degree_in_interf with (K := k). +assumption. +intro. elim (VertexSet.remove_1 (Register.eq_sym _ _ H4) H). +apply (proj1 (In_spill_props _ _ _ _ _ _ _ _ H1 (refl_equal _) HWS)). +apply (proj2 (In_spill_props _ _ _ _ _ _ _ _ H1 (refl_equal _) HWS)). +apply eq_K_1. +apply degree_dec_remove_K with (r := r). +intro. elim (VertexSet.remove_1 (Register.eq_sym _ _ H4) H). +apply (proj1 (In_spill_props _ _ _ _ _ _ _ _ H1 (refl_equal _) HWS)). +assumption. +reflexivity. +split. +unfold g'. rewrite In_remove_vertex. split. +apply (proj2 (In_spill_props _ _ _ _ _ _ _ _ H1 (refl_equal _) HWS)). +intro. elim (VertexSet.remove_1 (Register.eq_sym _ _ H3) H). +unfold g'. rewrite precolored_remove_vertex. +intro. elim (proj2 (proj2 (In_spill_props _ _ _ _ _ _ _ _ H1 (refl_equal _) HWS))). +apply (VertexSet.remove_3 H3). + +(* spillWL' <= *) +destruct H. destruct H0. +unfold get_spillWL. simpl. +unfold spillWL'. +assert (~Register.eq r x) as Hr. +intro. rewrite <-H2 in H0. unfold g' in H0. +rewrite In_remove_vertex in H0. destruct H0. elim H3. auto. +apply VertexSet.remove_2. assumption. +unfold spillWL_. apply VertexSet.diff_3. +WS_apply HWS. +split. +apply not_low_remove_not_low with (r:=r). +assumption. +intuition. +split. +unfold g' in H0. rewrite In_remove_vertex in H0. intuition. +unfold g' in H1. rewrite precolored_remove_vertex in H1. +intro. elim H1. apply VertexSet.remove_2; auto. +intro. unfold low in H2. +generalize (VertexSet.filter_1 (eq_K_compat k g) H2). intro H3. +generalize (VertexSet.filter_2 (eq_K_compat k g) H2). clear H2. intro H2. +generalize (eq_K_2 _ _ H2). clear H2. intro H2. +assert (has_low_degree g' k x = true). +apply degree_K_remove_dec. intuition. +unfold has_low_degree, interf_degree. rewrite <-H2. +destruct (le_lt_dec k k). reflexivity. elim (lt_irrefl _ l). +auto. +unfold not_pre_int_adj in H3. generalize (VertexSet.diff_1 H3). intro. +unfold int_adj in H4. assumption. +rewrite H in H4. inversion H4. + +(* freezeWL' => *) +split. +unfold get_freezeWL. simpl. split; intros. +unfold freezeWL' in H. +assert (~Register.eq r x) as Hr. +intro. elim (VertexSet.remove_1 H0 H). +generalize (VertexSet.remove_3 H). clear H. intro. +unfold freezeWL_ in H. +destruct (VertexSet.union_1 H). +unfold freezeWL__ in H0. +generalize (VertexSet.diff_1 H0). intro. +generalize (VertexSet.diff_2 H0). clear H0. intro. +split. +apply low_remove_low. +apply (proj1 (In_freeze_props _ _ _ _ _ _ _ _ H1 (refl_equal _) HWS)). +intuition. +split. +unfold nmr in H0. +case_eq (move_related g' x); intros. +reflexivity. +elim H0. apply VertexSet.filter_3. +apply compat_move_up. +unfold not_pre_pre_adj. apply VertexSet.diff_3. +unfold pre_adj. apply move_related_not_remove_in_pref. intuition. +apply (proj1 (proj2 (In_freeze_props _ _ _ _ _ _ _ _ H1 (refl_equal _) HWS))). +assumption. +apply (proj2 (proj2 (In_freeze_props _ _ _ _ _ _ _ _ H1 (refl_equal _) HWS))). +assert (VertexSet.cardinal (preference_adj x g) = 1). +apply pref_degree_dec_remove_1 with (r:=r). intuition. +apply (proj1 (proj2 (In_freeze_props _ _ _ _ _ _ _ _ H1 (refl_equal _) HWS))). assumption. +rewrite (eq_K_1 _ _ H3). simpl. +apply (proj1 (In_freeze_props _ _ _ _ _ _ _ _ H1 (refl_equal _) HWS)). +unfold g'. rewrite precolored_remove_vertex. +intro. elim (proj2 (proj2 (proj2 (In_freeze_props _ _ _ _ _ _ _ _ H1 (refl_equal _) HWS)))). +apply (VertexSet.remove_3 H2). +split. +assert (VertexSet.cardinal (interference_adj x g) = k). +assert (VertexSet.In x (fst (VertexSet.partition (move_related g) low))). +rewrite Hsf. simpl. assumption. +rewrite (VertexSet.partition_1) in H1. +generalize (VertexSet.filter_1 (compat_bool_move _) H1). intro. +unfold low in H2. +generalize (VertexSet.filter_1 (eq_K_compat k g) H2). intro H3. +generalize (VertexSet.filter_2 (eq_K_compat k g) H2). clear H2. intro H2. +generalize (eq_K_2 _ _ H2). auto. apply compat_bool_move. +apply degree_K_remove_dec. +intuition. +unfold has_low_degree, interf_degree. rewrite H1. +destruct (le_lt_dec k k). reflexivity. elim (lt_irrefl _ l). auto. +assert (VertexSet.In x (fst (VertexSet.partition (move_related g) low))). +rewrite Hsf. simpl. assumption. +rewrite (VertexSet.partition_1) in H2. +generalize (VertexSet.filter_1 (compat_bool_move _) H2). intro. +unfold low in H3. +generalize (VertexSet.filter_1 (eq_K_compat k g) H3). intro H4. +generalize (VertexSet.filter_2 (eq_K_compat k g) H3). clear H3. intro H3. +unfold not_pre_int_adj in H4. apply (VertexSet.diff_1 H4). +apply compat_bool_move. +split. +case_eq (move_related g' x); intros. +reflexivity. +assert (VertexSet.In x (fst (VertexSet.partition (move_related g) low))). +rewrite Hsf. simpl. assumption. +rewrite (VertexSet.partition_1) in H2. +generalize (VertexSet.filter_1 (compat_bool_move _) H2). intro. +generalize (VertexSet.filter_2 (compat_bool_move _) H2). clear H2. intro H2. +assert (VertexSet.In x (preference_adj r g)). +apply move_related_not_remove_in_pref; assumption. +assert (VertexSet.In x (interference_adj r g)). +unfold low in H3. +generalize (VertexSet.filter_1 (eq_K_compat k g) H3). intro H5. +apply (VertexSet.diff_1 H5). +elim (interf_pref_conflict x r g). split. +rewrite in_pref in H4. destruct H4. +unfold Prefere. exists x0. assumption. +rewrite in_interf in H5. unfold Interfere. assumption. +apply compat_bool_move. +unfold g'. rewrite precolored_remove_vertex. +intro. +assert (VertexSet.In x (fst (VertexSet.partition (move_related g) low))). +rewrite Hsf. simpl. assumption. +rewrite (VertexSet.partition_1) in H2. +generalize (VertexSet.filter_1 (compat_bool_move _) H2). intro. +unfold low in H3. +generalize (VertexSet.filter_1 (eq_K_compat k g) H3). intro. +elim (VertexSet.diff_2 H4). apply (VertexSet.remove_3 H1). +apply compat_bool_move. + +(* freezeWL' <= *) +destruct H. destruct H0. +unfold freezeWL'. +assert (~Register.eq r x). +intro. generalize (move_related_in_graph _ _ H0). intro. rewrite <-H2 in H3. +unfold g' in H3. rewrite In_remove_vertex in H3. destruct H3. elim H4. auto. +apply VertexSet.remove_2. assumption. +unfold freezeWL_. +case_eq (has_low_degree g k x); intros. +apply VertexSet.union_2. unfold freezeWL__. +apply VertexSet.diff_3. +WS_apply HWS. +split. +assumption. +split. +apply move_remove_2 with (r:=r). assumption. +unfold g' in H1. rewrite precolored_remove_vertex in H1. +intro. elim H1. apply VertexSet.remove_2. assumption. assumption. +intro. unfold nmr in H4. +assert (move_related g' x = false). +apply pref_degree_1_remove_dec. +intuition. +generalize (VertexSet.filter_2 (compat_move_up g k) H4). intro. +case_eq (eq_K 1 (VertexSet.cardinal (preference_adj x g))); intros. +rewrite (eq_K_2 _ _ H6); auto. +rewrite H6 in H5. inversion H5. +generalize (VertexSet.filter_1 (compat_move_up g k) H4). intro. +apply (VertexSet.diff_1 H5). +rewrite H0 in H5. inversion H5. +apply VertexSet.union_3. +cut (VertexSet.In x (fst (VertexSet.partition (move_related g) low))). intro. +rewrite Hsf in H4. simpl in H4. assumption. +rewrite VertexSet.partition_1. +apply VertexSet.filter_3. +apply compat_bool_move. +unfold low. apply VertexSet.filter_3. +apply eq_K_compat. +apply VertexSet.diff_3. +unfold int_adj. +apply low_degree_in_interf with (K:=k). +assumption. intuition. assumption. +unfold g' in H1. rewrite precolored_remove_vertex in H1. +intro. elim H1. apply VertexSet.remove_2; auto. +apply eq_K_1. +apply degree_dec_remove_K with (r:=r). +intuition. assumption. assumption. +apply move_remove_2 with (r:=r). assumption. +apply compat_bool_move. + +(* simplifyWL' => *) +split. +unfold get_simplifyWL. simpl. +split; intros. +unfold simplifyWL' in H. +assert (~Register.eq r x) as Hr. +intro. elim (VertexSet.remove_1 H0 H). +generalize (VertexSet.remove_3 H). clear H. intro H. +unfold simplifyWL_ in H. +destruct (VertexSet.union_1 H). +unfold simplifyWL__ in H0. +destruct (VertexSet.union_1 H0). +generalize (In_simplify_props _ _ _ _ _ _ _ _ H1 (refl_equal _) HWS). clear H1. intro. +destruct H1. destruct H2. destruct H3. +split. apply low_remove_low. assumption. intuition. +split. unfold g'. +apply move_remove. assumption. +split. unfold g'. rewrite In_remove_vertex. split; auto. +unfold g'. rewrite precolored_remove_vertex. intro. +elim H4. apply (VertexSet.remove_3 H5). +assert (VertexSet.In x (snd (VertexSet.partition (move_related g) low))). +rewrite Hsf. simpl. assumption. +rewrite VertexSet.partition_2 in H2. +generalize (VertexSet.filter_1 (compat_not_compat (compat_bool_move g)) H2). intro. +generalize (VertexSet.filter_2 (compat_not_compat (compat_bool_move g)) H2). clear H2. intro. +unfold low in H3. +generalize (VertexSet.filter_1 (eq_K_compat k g) H3). intro. +generalize (VertexSet.filter_2 (eq_K_compat k g) H3). clear H3. intro. +generalize (eq_K_2 _ _ H3). clear H3. intro. +split. +apply degree_K_remove_dec. intuition. +unfold has_low_degree, interf_degree. rewrite <-H3. +destruct (le_lt_dec k k). auto. elim (lt_irrefl _ l). auto. +apply (VertexSet.diff_1 H4). +split. +apply move_remove. +destruct (move_related g x). inversion H2. reflexivity. +split. +unfold g'. rewrite In_remove_vertex. split. +apply in_interf_in with (r:=r). +apply (VertexSet.diff_1 H4). intuition. +unfold g'. rewrite precolored_remove_vertex. intro. +elim (VertexSet.diff_2 H4). apply (VertexSet.remove_3 H5). +apply compat_bool_move. +(* x in nmr *) +unfold nmr in H0. +generalize (VertexSet.filter_1 (compat_move_up g k) H0). intro. +generalize (VertexSet.filter_2 (compat_move_up g k) H0). clear H0. intro. +assert (VertexSet.cardinal (preference_adj x g) = 1). +symmetry. apply eq_K_2. +case_eq (eq_K 1 (VertexSet.cardinal (preference_adj x g))); intros. +reflexivity. +rewrite H2 in H0. inversion H0. +rewrite (eq_K_1 _ _ H2) in H0. simpl in H0. +split. +apply low_remove_low. assumption. intuition. +split. +case_eq (move_related g x); intros. +apply pref_degree_1_remove_dec. intuition. assumption. +apply (VertexSet.diff_1 H1). +apply move_remove. assumption. +split. unfold g'. rewrite In_remove_vertex. split. apply in_pref_in with (r:=r). +apply (VertexSet.diff_1 H1). intuition. +unfold g'. rewrite precolored_remove_vertex. intro. +elim (VertexSet.diff_2 H1). apply (VertexSet.remove_3 H3). + +(* simplifyWL' <= *) +destruct H. destruct H0. destruct H1. +unfold simplifyWL'. +assert (~Register.eq r x) as Hr. +intro. rewrite <-H3 in H1. +unfold g' in H1. rewrite In_remove_vertex in H1. destruct H1. elim H4. auto. +apply VertexSet.remove_2. assumption. +unfold simplifyWL_. +case_eq (has_low_degree g k x); intros. +case_eq (move_related g x); intros. +apply VertexSet.union_3. +unfold nmr. apply VertexSet.filter_3. +apply compat_move_up. +apply VertexSet.diff_3. +apply move_related_not_remove_in_pref. assumption. assumption. assumption. +unfold g' in H2. rewrite precolored_remove_vertex in H2. +intro. elim H2. apply VertexSet.remove_2. intuition. assumption. +assert (eq_K 1 (VertexSet.cardinal (preference_adj x g)) = true). +apply eq_K_1. +apply pref_degree_dec_remove_1 with (r:=r). intuition . assumption. assumption. +rewrite H5. simpl. assumption. +apply VertexSet.union_2. +unfold simplifyWL__. +apply VertexSet.union_2. +WS_apply HWS. +split. +assumption. +split. +assumption. +split. +unfold g' in H1. rewrite In_remove_vertex in H1. intuition. +unfold g' in H2. rewrite precolored_remove_vertex in H2. +intro. elim H2. apply VertexSet.remove_2. intuition. assumption. +apply VertexSet.union_2. +unfold simplifyWL__. +apply VertexSet.union_3. +assert (VertexSet.In x (snd (VertexSet.partition (move_related g) low))). +rewrite VertexSet.partition_2. +apply VertexSet.filter_3. +apply compat_not_compat. apply compat_bool_move. +unfold low. +apply VertexSet.filter_3. +apply eq_K_compat. apply VertexSet.diff_3. +apply low_degree_in_interf with (K:=k). assumption. intuition. assumption. +unfold g' in H2. rewrite precolored_remove_vertex in H2. +intro. elim H2. apply VertexSet.remove_2. intuition. assumption. +apply eq_K_1. apply degree_dec_remove_K with (r:=r). intuition. assumption. assumption. +case_eq (move_related g x); intros. +assert (VertexSet.In x (interference_adj r g)). +apply low_degree_in_interf with (K:=k). assumption. intuition. assumption. +assert (VertexSet.In x (preference_adj r g)). +apply move_related_not_remove_in_pref. assumption. assumption. assumption. +elim (interf_pref_conflict x r g). +split. +unfold Prefere. rewrite in_pref in H6. assumption. +unfold Interfere. rewrite <-in_interf. assumption. +auto. +apply compat_bool_move. +rewrite Hsf in H4. simpl in H4. assumption. + +(* movesWL => *) +unfold movesWL', get_movesWL. simpl. +split; intros. +simpl in H. +rewrite not_incident_edges_1 in H. destruct H as [H HH]. +generalize (In_move_props _ _ _ _ _ _ _ _ H (refl_equal _) HWS). clear H. intro H. +destruct H. destruct H0. +split. assumption. +unfold g'. rewrite In_remove_edge. split. +apply (proj2 (proj1 (In_graph_aff_edge_in_AE _ _) H0)). +assumption. +generalize (proj1 (In_graph_interf_edge_in_IE _ _) H0). intro. +destruct H1. +inversion H. rewrite H1 in H3. inversion H3. +intros. apply (In_move_props _ _ _ _ _ _ _ _ H0 (refl_equal _) HWS). +destruct H. +rewrite not_incident_edges_1. +split. +WS_apply HWS. +split. +assumption. +unfold g' in H0. rewrite In_remove_edge in H0. intuition. +intro H1. +elim (In_graph_edge_in_ext _ _ H0). +intros. +destruct H1. +rewrite <-H1 in H2. +unfold g' in H2. rewrite In_remove_vertex in H2. destruct H2. elim H4. auto. +rewrite <-H1 in H3. +unfold g' in H3. rewrite In_remove_vertex in H3. destruct H3. elim H4. auto. +intros. apply (In_move_props _ _ _ _ _ _ _ _ H1 (refl_equal _) HWS). +Qed. diff --git a/backend/SetsFacts.v b/backend/SetsFacts.v new file mode 100755 index 00000000..60145cf0 --- /dev/null +++ b/backend/SetsFacts.v @@ -0,0 +1,105 @@ +Require Import FSets. +Set Implicit Arguments. + +(* Useful properties over finite sets and ordered types *) + +Module MyOTFacts (M : OrderedType). + +Section compat. + +Lemma compat_not_compat : forall f : M.t -> bool, +compat_bool M.eq f -> +compat_bool M.eq (fun x => negb (f x)). + +Proof. +unfold compat_bool;intros f H. +intros; unfold negb. +rewrite (H x y H0);reflexivity. +Qed. + +End compat. + +End MyOTFacts. + +Module MyFacts (M : S). + +Import M. +Module Import Props := Properties M. +Module Import Facts := OrderedTypeFacts E. +Module Import OTFacts := MyOTFacts E. + +Lemma set_induction2 : forall s, +Empty s \/ exists x, exists s', Add x s' s. + +Proof. +intros. case_eq (choose s);intros. +right. exists e. exists (remove e s). +constructor;intro H0. +destruct (eq_dec e y). +left;assumption. +right;apply (remove_2 n H0). +destruct H0. +rewrite <-H0;apply (choose_1 H). +eapply remove_3;eassumption. +left;apply (choose_2 H). +Qed. + +Lemma equal_equivlist : forall s s', +Equal s s' -> equivlistA E.eq (elements s) (elements s'). + +Proof. +unfold equivlistA. +generalize elements_1;generalize elements_2;intros H0 H1 s s' H x. +split;intro H2. +apply H1;rewrite <-H;apply H0;assumption. +apply H1;rewrite H;apply H0;assumption. +Qed. + +Section Fold_Facts. + +Variable A : Type. + +Lemma fold_left_compat_set : forall (f : t -> A -> t) l e e', +Equal e e' -> +(forall e1 e2 a, Equal e1 e2 -> Equal (f e1 a) (f e2 a)) -> +Equal (fold_left f l e) (fold_left f l e'). + +Proof. +intros f l. +induction l;simpl. +auto. +intros e e' H H0 H1. +apply (IHl (f e a) (f e' a)). +apply H0;assumption. +assumption. +Qed. + +Lemma fold_left_assoc : forall l f x h, +(forall (y z : A) s, Equal (f (f s y) z) (f (f s z) y)) -> +(forall e1 e2 a, Equal e1 e2 -> Equal (f e1 a) (f e2 a)) -> +Equal (fold_left f (h :: l) x) (f (fold_left f l x) h). + +Proof. +induction l;simpl;intros f x h H H0. +intuition. +rewrite <-IHl;simpl;try assumption. +apply fold_left_compat_set;[apply H|];auto. +Qed. + +Lemma NoDupA_elements : forall s, +NoDupA E.eq (elements s). + +Proof. +intro s. +apply SortA_NoDupA with (ltA := E.lt). +apply E.eq_refl. +apply E.eq_sym. +apply E.lt_trans. +apply E.lt_not_eq. +apply Facts.lt_eq. +apply Facts.eq_lt. +apply elements_3. +Qed. + +End Fold_Facts. +End MyFacts. \ No newline at end of file diff --git a/backend/Simplify_WL.v b/backend/Simplify_WL.v new file mode 100755 index 00000000..4da207a9 --- /dev/null +++ b/backend/Simplify_WL.v @@ -0,0 +1,160 @@ +Require Import FSets. +Require Import InterfGraphMapImp. +Require Import Remove_Vertex_WL. +Require Import WS. +Require Import Affinity_relation. +Require Import Interference_adjacency. +Require Import Edges. +Require Import IRC_graph. + +Import Edge RegFacts Props OTFacts. + +Definition simplify_wl r ircg K := + let g := irc_g ircg in + let wl := irc_wl ircg in + let simplify := get_simplifyWL wl in + let freeze := get_freezeWL wl in + let spillWL := get_spillWL wl in + let movesWL := get_movesWL wl in + let pre := precolored g in + let int_adj := interference_adj r g in + let not_pre_int_adj := VertexSet.diff int_adj pre in + let newlow := VertexSet.filter (fun x => eq_K K (VertexSet.cardinal (interference_adj x g))) not_pre_int_adj in + let (free, simp) := VertexSet.partition (move_related g) newlow in + let simplifyWL_ := VertexSet.union simplify simp in + let simplifyWL' := VertexSet.remove r simplifyWL_ in + let freezeWL' := VertexSet.union freeze free in + let spillWL' := VertexSet.diff spillWL newlow in + let movesWL' := movesWL in (spillWL', freezeWL', simplifyWL', movesWL'). + +Lemma WS_simplify_aux : forall r ircg, +VertexSet.In r (get_simplifyWL (irc_wl ircg)) -> +WS_properties (remove_vertex r (irc_g ircg)) (VertexSet.cardinal (pal ircg)) + (simplify_wl r ircg (VertexSet.cardinal (pal ircg))). + +Proof. +intros. +generalize (WS_props_equal (remove_vertex r (irc_g ircg)) (VertexSet.cardinal (pal ircg)) + (remove_wl_2 r ircg (VertexSet.cardinal (pal ircg))) + (simplify_wl r ircg (VertexSet.cardinal (pal ircg)))). +generalize (WS_remove_wl_2 r ircg). +unfold remove_wl_2, simplify_wl, get_simplifyWL, + get_freezeWL, get_spillWL, get_movesWL. +set (g' := remove_vertex r (irc_g ircg)) in *. +set (k := VertexSet.cardinal (pal ircg)) in *. +set (g := irc_g ircg) in *. +set (wl := irc_wl ircg) in *. +set ( simplify := get_simplifyWL wl ) in *. +set ( freeze := get_freezeWL wl ) in *. +set ( spillWL := get_spillWL wl ) in *. +set ( int_adj := interference_adj r g ) in *. +set ( not_pre_int_adj := VertexSet.diff int_adj (precolored g) ) in *. +set ( pre_adj := preference_adj r g ) in *. +set ( not_pre_pre_adj := VertexSet.diff pre_adj (precolored g) ) in *. +set ( low := VertexSet.filter (fun x => eq_K k (VertexSet.cardinal (interference_adj x g))) not_pre_int_adj ) in *. +set ( simpfree := VertexSet.partition (move_related g) low ) in *. +case_eq simpfree. intros free simp Hsf. +unfold simpfree in Hsf. +set ( nmr := VertexSet.filter (fun x => eq_K 1 (VertexSet.cardinal (preference_adj x g)) && has_low_degree g k x) not_pre_pre_adj) in *. +set ( simplifyWL__ := VertexSet.union simplify simp ) in *. +set ( simplifyWL_ := VertexSet.union simplifyWL__ nmr) in *. +set ( simplifyWL' := VertexSet.remove r simplifyWL_ ) in *. +set ( freezeWL__ := VertexSet.diff freeze nmr ) in *. +set ( freezeWL_ := VertexSet.union freezeWL__ free ) in *. +set ( freezeWL' := VertexSet.remove r freezeWL_ ) in *. +set ( spillWL_ := VertexSet.diff spillWL low ) in *. +set ( spillWL' := VertexSet.remove r spillWL_ ) in *. +set ( movesWL' := not_incident_edges r (get_movesWL wl) g) in *. simpl. +generalize (In_simplify_props _ _ _ _ _ _ _ _ H (refl_equal _) (HWS_irc ircg)). intro Hr. + +assert (VertexSet.Equal (preference_adj r g) VertexSet.empty) as Hempty_pre. +apply not_move_related_empty_pref. intuition. +assert (VertexSet.Equal nmr VertexSet.empty) as Hempty_nmr. +unfold nmr. +split; intros. +generalize (VertexSet.filter_1 (compat_move_up _ _) H0). intro. +unfold not_pre_pre_adj in H1. generalize (VertexSet.diff_1 H1). intro. +unfold pre_adj in H2. rewrite Hempty_pre in H2. elim (VertexSet.empty_1 H2). +elim (VertexSet.empty_1 H0). + +intros HWS H0. apply H0. + +(* simplify worklist is only decreasing from r and increasing from simp, + since nmr is empty, because preference adj r g is empty *) +split; intros. +apply VertexSet.remove_2. +intro. elim (VertexSet.remove_1 H2 H1). +destruct (VertexSet.union_1 (VertexSet.remove_3 H1)). +assumption. +rewrite Hempty_nmr in H2. elim (VertexSet.empty_1 H2). + +apply VertexSet.remove_2. +intro. elim (VertexSet.remove_1 H2 H1). +apply VertexSet.union_2. apply (VertexSet.remove_3 H1). + +(* r is not deleted from freezewl, since it is not move-related and + no vertex is removed from freeze, since free is empty, because + preference adj r g is empty *) + +set (s := VertexSet.union (VertexSet.diff (snd (fst (fst wl))) nmr) free). +split; intros. +unfold s in H1. destruct (VertexSet.union_1 (VertexSet.remove_3 H1)). +apply VertexSet.union_2. apply (VertexSet.diff_1 H2). +apply VertexSet.union_3. assumption. + +apply VertexSet.remove_2. intro. rewrite <-H2 in H1. clear H2. +unfold s in H1. +destruct (VertexSet.union_1 H1). +generalize (In_freeze_props _ _ _ _ _ _ _ _ H2 (refl_equal _) (HWS_irc ircg)). intro. +destruct H3. destruct H4. destruct Hr. destruct H7. +rewrite H7 in H4. inversion H4. +assert (free = fst (VertexSet.partition (move_related g) low)). +rewrite Hsf. auto. +rewrite H3 in H2. rewrite VertexSet.partition_1 in H2. +generalize (VertexSet.filter_2 (compat_bool_move _ ) H2). intro. +destruct Hr. destruct H6. unfold g in H4. rewrite H6 in H4. inversion H4. +apply compat_bool_move. + +unfold s. destruct (VertexSet.union_1 H1). +apply VertexSet.union_2. apply VertexSet.diff_3. +assumption. +intro. rewrite Hempty_nmr in H3. elim (VertexSet.empty_1 H3). +apply VertexSet.union_3. assumption. + +(* identically, r is not removed from spillwl, since it is not of high-degree, + but degrees remain the same even if r is not move related *) +set (s := VertexSet.diff (fst (fst (fst wl))) low). +split; intros. +apply (VertexSet.remove_3 H1). +apply VertexSet.remove_2. intro. rewrite <-H2 in H1. clear H2. +unfold s in H1. +generalize (VertexSet.diff_1 H1). intro. +generalize (In_spill_props _ _ _ _ _ _ _ _ H2 (refl_equal _) (HWS_irc ircg)). intro. +destruct H3. destruct Hr. rewrite H5 in H3. inversion H3. +assumption. + +(* there is no preference edge incident to r *) +split; intros. +rewrite not_incident_edges_1 in H1. destruct H1. assumption. +intros. apply (In_move_props _ _ _ _ _ _ _ _ H2 (refl_equal _) (HWS_irc ircg)). +rewrite not_incident_edges_1. split. +assumption. +intro. destruct Hr. destruct H4. +cut (move_related g r = true). intro. unfold g in H6. rewrite H4 in H6. inversion H6. +generalize (In_move_props _ _ _ _ _ _ _ _ H1 (refl_equal _) (HWS_irc ircg)). intro. +destruct H6. +apply move_related_charac2 with (e:=a); assumption. +intros. apply (In_move_props _ _ _ _ _ _ _ _ H2 (refl_equal _) (HWS_irc ircg)). + +(* WS_remove_wl respects the invariant *) +assumption. +Qed. + +Lemma WS_simplify : forall r ircg, +VertexSet.In r (get_simplifyWL (irc_wl ircg)) -> +WS_properties (remove_vertex r (irc_g ircg)) (irc_k ircg) + (simplify_wl r ircg (irc_k ircg)). + +Proof. +intros. rewrite <-(Hk ircg). apply WS_simplify_aux. auto. +Qed. diff --git a/backend/Spill_WL.v b/backend/Spill_WL.v new file mode 100755 index 00000000..1ed1c2d9 --- /dev/null +++ b/backend/Spill_WL.v @@ -0,0 +1,141 @@ +Require Import FSets. +Require Import InterfGraphMapImp. +Require Import Freeze_WL. +Require Import Edges. +Require Import WS. +Require Import Interference_adjacency. +Require Import Affinity_relation. +Require Import Remove_Vertex_WL. +Require Import IRC_graph. + +Import RegFacts Props OTFacts. + +Definition spill_wl r ircg K := + let g := irc_g ircg in + let wl := irc_wl ircg in + let simplify := get_simplifyWL wl in + let freeze := get_freezeWL wl in + let spillWL := get_spillWL wl in + let movesWL := get_movesWL wl in + let pre := precolored g in + let int_adj := interference_adj r g in + let not_pre_int_adj := VertexSet.diff int_adj pre in + let pre_adj := preference_adj r g in + let not_pre_pre_adj := VertexSet.diff pre_adj pre in + let newlow := VertexSet.filter (fun x => eq_K K (VertexSet.cardinal (interference_adj x g))) not_pre_int_adj in + let (free, simp) := VertexSet.partition (move_related g) newlow in + let newnmr := VertexSet.filter (fun x => eq_K 1 (VertexSet.cardinal (preference_adj x g)) && has_low_degree g K x) not_pre_pre_adj in + let simplifyWL__ := VertexSet.union simplify simp in + let simplifyWL' := VertexSet.union simplifyWL__ newnmr in + let freezeWL__ := VertexSet.diff freeze newnmr in + let freezeWL' := VertexSet.union freezeWL__ free in + let spillWL_ := VertexSet.diff spillWL newlow in + let spillWL' := VertexSet.remove r spillWL_ in + let movesWL' := not_incident_edges r movesWL g in + (spillWL', freezeWL', simplifyWL', movesWL'). + + +Lemma WS_spill_aux : forall r ircg, +VertexSet.In r (get_spillWL (irc_wl ircg)) -> +WS_properties (remove_vertex r (irc_g ircg)) (VertexSet.cardinal (pal ircg)) + (spill_wl r ircg (VertexSet.cardinal (pal ircg))). + +Proof. +intros. +generalize (WS_props_equal (remove_vertex r (irc_g ircg)) (VertexSet.cardinal (pal ircg)) + (remove_wl_2 r ircg (VertexSet.cardinal (pal ircg))) + (spill_wl r ircg (VertexSet.cardinal (pal ircg)))). +generalize (WS_remove_wl_2 r ircg). +unfold remove_wl_2, spill_wl, get_simplifyWL, + get_freezeWL, get_spillWL, get_movesWL. +set (g' := remove_vertex r (irc_g ircg)) in *. +set (k := VertexSet.cardinal (pal ircg)) in *. +set (g := irc_g ircg) in *. +set (wl := irc_wl ircg) in *. +set ( simplify := get_simplifyWL wl ) in *. +set ( freeze := get_freezeWL wl ) in *. +set ( spillWL := get_spillWL wl ) in *. +set ( int_adj := interference_adj r g ) in *. +set ( not_pre_int_adj := VertexSet.diff int_adj (precolored g) ) in *. +set ( pre_adj := preference_adj r g ) in *. +set ( not_pre_pre_adj := VertexSet.diff pre_adj (precolored g) ) in *. +set ( low := VertexSet.filter (fun x => eq_K k (VertexSet.cardinal (interference_adj x g))) not_pre_int_adj ) in *. +set ( simpfree := VertexSet.partition (move_related g) low ) in *. +case_eq simpfree. intros free simp Hsf. +unfold simpfree in Hsf. +set ( nmr := VertexSet.filter (fun x => eq_K 1 (VertexSet.cardinal (preference_adj x g)) && has_low_degree g k x) not_pre_pre_adj) in *. +set ( simplifyWL__ := VertexSet.union simplify simp ) in *. +set ( simplifyWL_ := VertexSet.union simplifyWL__ nmr) in *. +set ( simplifyWL' := VertexSet.remove r simplifyWL_ ) in *. +set ( freezeWL__ := VertexSet.diff freeze nmr ) in *. +set ( freezeWL_ := VertexSet.union freezeWL__ free ) in *. +set ( freezeWL' := VertexSet.remove r freezeWL_ ) in *. +set ( spillWL_ := VertexSet.diff spillWL low ) in *. +set ( spillWL' := VertexSet.remove r spillWL_ ) in *. +set ( movesWL' := not_incident_edges r (get_movesWL wl) g) in *. simpl. +generalize (In_spill_props _ _ _ _ _ _ _ _ H (refl_equal _) (HWS_irc ircg)). intro Hr. + +intros HWS H0. apply H0. + +(* r is not removed from simplify *) +split; intros. +apply (VertexSet.remove_3 H1). + +apply VertexSet.remove_2. intro. rewrite <-H2 in H1. clear H2. +destruct (VertexSet.union_1 H1). +destruct (VertexSet.union_1 H2). +generalize (In_simplify_props _ _ _ _ _ _ _ _ H3 (refl_equal _) (HWS_irc ircg)). intro. +destruct H4. destruct Hr. rewrite H4 in H6. inversion H6. + +assert (simp = snd (VertexSet.partition (move_related g) low)) as Hsimp. +rewrite Hsf. auto. rewrite Hsimp in H3. +rewrite VertexSet.partition_2 in H3. +generalize (VertexSet.filter_1 (compat_not_compat (compat_bool_move _)) H3). intro. +unfold low in H4. generalize (VertexSet.filter_1 (eq_K_compat _ _) H4). intro. +elim (not_in_interf_self r g). apply (VertexSet.diff_1 H5). +apply compat_bool_move. + +unfold nmr in H2. generalize (VertexSet.filter_1 (compat_move_up _ _) H2). intro. +elim (not_in_pref_self r g (VertexSet.diff_1 H3)). +assumption. + +(* r is not deleted from freezewl, since it is not move-related and + no vertex is removed from freeze, since free is empty, because + preference adj r g is empty *) + +set (s := VertexSet.union (VertexSet.diff (snd (fst (fst wl))) nmr) free). +split; intros. +apply (VertexSet.remove_3 H1). + +apply VertexSet.remove_2. intro. rewrite <-H2 in H1. clear H2. +unfold s in H1. +destruct (VertexSet.union_1 H1). +generalize (In_freeze_props _ _ _ _ _ _ _ _ (VertexSet.diff_1 H2) (refl_equal _) (HWS_irc ircg)). intro. +destruct H3. destruct Hr. rewrite H5 in H3. inversion H3. + +assert (free = fst (VertexSet.partition (move_related g) low)). +rewrite Hsf. auto. +rewrite H3 in H2. rewrite VertexSet.partition_1 in H2. +generalize (VertexSet.filter_1 (compat_bool_move _ ) H2). intro. +unfold low in H4. generalize (VertexSet.filter_1 (eq_K_compat _ _) H4). intro. +elim (not_in_interf_self r g). apply (VertexSet.diff_1 H5). +apply compat_bool_move. assumption. + +(* spill worklist is unchanged *) +apply VertexSet.eq_refl. + +(* moves worklst is unchanged *) +apply EdgeSet.eq_refl. + +(* WS_remove_wl respects the invariant *) +assumption. +Qed. + +Lemma WS_spill : forall r ircg, +VertexSet.In r (get_spillWL (irc_wl ircg)) -> +WS_properties (remove_vertex r (irc_g ircg)) (irc_k ircg) + (spill_wl r ircg (irc_k ircg)). + +Proof. +intros. rewrite <-(Hk ircg). apply WS_spill_aux. auto. +Qed. diff --git a/backend/Typed_interfgraphs.v b/backend/Typed_interfgraphs.v new file mode 100755 index 00000000..0f252288 --- /dev/null +++ b/backend/Typed_interfgraphs.v @@ -0,0 +1,740 @@ +Require Import InterfGraph. +Require Import AST. +Require Import FSets. +Require Import Locations. +Require Import Registers. + +Module Import SRRFacts := Facts SetRegReg. +Module MRegset := FSetAVL.Make OrderedMreg. + +Close Scope nat_scope. + +Definition regregpartition : Type := SetRegReg.t*SetRegReg.t*Regset.t*Regset.t. + +Definition rr1 := fun (p : regregpartition) => fst (fst (fst p)). +Definition rr2 := fun (p : regregpartition) => snd (fst (fst p)). +Definition rr3 := fun (p : regregpartition) => snd (fst p). +Definition rr4 := fun (p : regregpartition) => snd p. + +Definition regreg_edge_type_partition s env := +SetRegReg.fold (fun e s => match env (fst e), env (snd e) with + | Tint, Tint => (SetRegReg.add e (rr1 s), rr2 s, + Regset.add (fst e) (Regset.add (snd e) (rr3 s)), rr4 s) + | Tfloat, Tfloat => (rr1 s, SetRegReg.add e (rr2 s), rr3 s, + Regset.add (fst e) (Regset.add (snd e) (rr4 s))) + | Tint, Tfloat => (rr1 s, rr2 s, Regset.add (fst e) (rr3 s), Regset.add (snd e) (rr4 s)) + | Tfloat, Tint => (rr1 s, rr2 s, Regset.add (snd e) (rr3 s), Regset.add (fst e) (rr4 s)) + end) + s + (SetRegReg.empty, SetRegReg.empty, Regset.empty, Regset.empty). + +Lemma in_partition_in_fst : forall e s env, +SetRegReg.In e (rr1 (regreg_edge_type_partition s env)) -> +SetRegReg.In e s. + +Proof. +intros e s env H. +unfold regreg_edge_type_partition in H. +set (f := fun (e : SetRegReg.elt) (s : regregpartition) => + match env (fst e) with + | Tint => + match env (snd e) with + | Tint => + (SetRegReg.add e (rr1 s), rr2 s, + Regset.add (fst e) (Regset.add (snd e) (rr3 s)), + rr4 s) + | Tfloat => + (rr1 s, rr2 s, Regset.add (fst e) (rr3 s), + Regset.add (snd e) (rr4 s)) + end + | Tfloat => + match env (snd e) with + | Tint => + (rr1 s, rr2 s, Regset.add (snd e) (rr3 s), + Regset.add (fst e) (rr4 s)) + | Tfloat => + (rr1 s, SetRegReg.add e (rr2 s), rr3 s, + Regset.add (fst e) (Regset.add (snd e) (rr4 s))) + end + end). +unfold regregpartition in *. fold f in H. + +assert (forall e set1 set2 set3 set4, SetRegReg.In e (rr1 (SetRegReg.fold f s (set1, set2, set3, set4))) -> + SetRegReg.In e s \/ SetRegReg.In e set1 \/ SetRegReg.In e set2). +clear H. +intros e' s1 s2 s3 s4 H. +rewrite SetRegReg.fold_1 in H. +generalize H. generalize s1 s2 s3 s4. clear s1 s2 s3 s4 H. +generalize (SetRegReg.elements_2). intro HH. +generalize (HH s). clear HH. intro HH. +induction (SetRegReg.elements s). +simpl. right. left. assumption. +intros s1 s2 s3 s4 H. +simpl in H. +assert ((forall x : SetRegReg.elt, + SetoidList.InA (fun x0 y : OrderedRegReg.t => fst x0 = fst y /\ snd x0 = snd y) x l -> + SetRegReg.In x s)). +intros. apply HH. right. assumption. +generalize (IHl H0). clear IHl H0. intro IHl. +assert (f a (s1, s2, s3, s4) = (SetRegReg.add a s1, s2, Regset.add (fst a) (Regset.add (snd a) s3), s4) \/ + f a (s1, s2, s3, s4) = (s1, SetRegReg.add a s2, s3, Regset.add (fst a) (Regset.add (snd a) s4)) \/ + f a (s1, s2, s3, s4) = (s1, s2, Regset.add (fst a) s3, Regset.add (snd a) s4) \/ + f a (s1, s2, s3, s4) = (s1, s2, Regset.add (snd a) s3, Regset.add (fst a) s4)). +unfold f. +destruct (env (snd a)); destruct (env (fst a)); unfold rr1, rr2, rr3, rr4; simpl. +left. reflexivity. +right. right. right. reflexivity. +right. right. left. reflexivity. +right. left. reflexivity. +destruct H0. +rewrite H0 in H. + +generalize (IHl (SetRegReg.add a s1) s2 _ _ H). +intro H1. destruct H1. +left. assumption. +destruct H1. + +destruct (proj1 (add_iff _ _ _) H1). +left. apply HH. left. intuition. +right. left. assumption. +right. right. assumption. +destruct H0. +rewrite H0 in H. +generalize (IHl s1 (SetRegReg.add a s2) _ _ H). +intro H1. destruct H1. +left. assumption. +destruct H1. +right. left. assumption. +destruct (proj1 (add_iff _ _ _) H1). +left. apply HH. left. intuition. +right. right. assumption. +destruct H0. +rewrite H0 in H. +generalize (IHl s1 s2 _ _ H). +intro H1. destruct H1. +left. assumption. +destruct H1. +right. left. assumption. +right. right. assumption. +rewrite H0 in H. +generalize (IHl s1 s2 _ _ H). +intro H1. destruct H1. +left. assumption. +destruct H1. +right. left. assumption. +right. right. assumption. +generalize (H0 _ _ _ _ _ H). clear H0. intro H0. +destruct H0. +assumption. +destruct H0; elim (SetRegReg.empty_1 H0). +Qed. + +Lemma in_partition_in_snd : forall e s env, +SetRegReg.In e (rr2 (regreg_edge_type_partition s env)) -> +SetRegReg.In e s. + +Proof. +intros e s env H. +unfold regreg_edge_type_partition in H. +set (f := fun (e : SetRegReg.elt) (s : regregpartition) => + match env (fst e) with + | Tint => + match env (snd e) with + | Tint => + (SetRegReg.add e (rr1 s), rr2 s, + Regset.add (fst e) (Regset.add (snd e) (rr3 s)), + rr4 s) + | Tfloat => + (rr1 s, rr2 s, Regset.add (fst e) (rr3 s), + Regset.add (snd e) (rr4 s)) + end + | Tfloat => + match env (snd e) with + | Tint => + (rr1 s, rr2 s, Regset.add (snd e) (rr3 s), + Regset.add (fst e) (rr4 s)) + | Tfloat => + (rr1 s, SetRegReg.add e (rr2 s), rr3 s, + Regset.add (fst e) (Regset.add (snd e) (rr4 s))) + end + end). +unfold regregpartition in *. fold f in H. + +assert (forall e set1 set2 set3 set4, SetRegReg.In e (rr2 (SetRegReg.fold f s (set1, set2, set3, set4))) -> + SetRegReg.In e s \/ SetRegReg.In e set1 \/ SetRegReg.In e set2). +clear H. +intros e' s1 s2 s3 s4 H. +rewrite SetRegReg.fold_1 in H. +generalize H. generalize s1 s2 s3 s4. clear s1 s2 s3 s4 H. +generalize (SetRegReg.elements_2). intro HH. +generalize (HH s). clear HH. intro HH. +induction (SetRegReg.elements s). +simpl. right. right. assumption. +intros s1 s2 s3 s4 H. +simpl in H. +assert ((forall x : SetRegReg.elt, + SetoidList.InA (fun x0 y : OrderedRegReg.t => fst x0 = fst y /\ snd x0 = snd y) x l -> + SetRegReg.In x s)). +intros. apply HH. right. assumption. +generalize (IHl H0). clear IHl H0. intro IHl. +assert (f a (s1, s2, s3, s4) = (SetRegReg.add a s1, s2, Regset.add (fst a) (Regset.add (snd a) s3), s4) \/ + f a (s1, s2, s3, s4) = (s1, SetRegReg.add a s2, s3, Regset.add (fst a) (Regset.add (snd a) s4)) \/ + f a (s1, s2, s3, s4) = (s1, s2, Regset.add (fst a) s3, Regset.add (snd a) s4) \/ + f a (s1, s2, s3, s4) = (s1, s2, Regset.add (snd a) s3, Regset.add (fst a) s4)). +unfold f. +destruct (env (snd a)); destruct (env (fst a)); unfold rr1, rr2, rr3, rr4; simpl. +left. reflexivity. +right. right. right. reflexivity. +right. right. left. reflexivity. +right. left. reflexivity. +destruct H0. +rewrite H0 in H. + +generalize (IHl (SetRegReg.add a s1) s2 _ _ H). +intro H1. destruct H1. +left. assumption. +destruct H1. + +destruct (proj1 (add_iff _ _ _) H1). +left. apply HH. left. intuition. +right. left. assumption. +right. right. assumption. +destruct H0. +rewrite H0 in H. +generalize (IHl s1 (SetRegReg.add a s2) _ _ H). +intro H1. destruct H1. +left. assumption. +destruct H1. +right. left. assumption. +destruct (proj1 (add_iff _ _ _) H1). +left. apply HH. left. intuition. +right. right. assumption. +destruct H0. +rewrite H0 in H. +generalize (IHl s1 s2 _ _ H). +intro H1. destruct H1. +left. assumption. +destruct H1. +right. left. assumption. +right. right. assumption. +rewrite H0 in H. +generalize (IHl s1 s2 _ _ H). +intro H1. destruct H1. +left. assumption. +destruct H1. +right. left. assumption. +right. right. assumption. +generalize (H0 _ _ _ _ _ H). clear H0. intro H0. +destruct H0. +assumption. +destruct H0; elim (SetRegReg.empty_1 H0). +Qed. + +Lemma in_partition_type_fst : forall e s env, +SetRegReg.In e (rr1 (regreg_edge_type_partition s env)) -> +env (fst e) = Tint /\ env (snd e) = Tint. + +Proof. +intros e s env H. +unfold regreg_edge_type_partition in H. +set (f := fun (e : SetRegReg.elt) (s : regregpartition) => + match env (fst e) with + | Tint => + match env (snd e) with + | Tint => + (SetRegReg.add e (rr1 s), rr2 s, + Regset.add (fst e) (Regset.add (snd e) (rr3 s)), + rr4 s) + | Tfloat => + (rr1 s, rr2 s, Regset.add (fst e) (rr3 s), + Regset.add (snd e) (rr4 s)) + end + | Tfloat => + match env (snd e) with + | Tint => + (rr1 s, rr2 s, Regset.add (snd e) (rr3 s), + Regset.add (fst e) (rr4 s)) + | Tfloat => + (rr1 s, SetRegReg.add e (rr2 s), rr3 s, + Regset.add (fst e) (Regset.add (snd e) (rr4 s))) + end + end). +unfold regregpartition in *. fold f in H. + +cut (forall e set1 set2 set3 set4, SetRegReg.In e (rr1 (SetRegReg.fold f s (set1, set2, set3, set4))) -> + (env (fst e) = Tint /\ env (snd e) = Tint) \/ SetRegReg.In e set1). + +intros. +generalize (H0 _ _ _ _ _ H). clear H H0. intro H. +destruct H. assumption. elim (SetRegReg.empty_1 H). + +(* cut property *) +intros. +generalize H0. clear H H0. intro H. +rewrite SetRegReg.fold_1 in H. +generalize H. clear H. generalize set1 set2 set3 set4. +induction (SetRegReg.elements s); intros s1 s2 s3 s4 H. +simpl in H. right. assumption. +simpl in H. + +assert ((f a (s1, s2, s3, s4) = (SetRegReg.add a s1, s2, Regset.add (fst a) (Regset.add (snd a) s3), s4) /\ + env (fst a) = Tint /\ + env (snd a) = Tint)\/ + (f a (s1, s2, s3, s4) = (s1, SetRegReg.add a s2, s3, Regset.add (fst a) (Regset.add (snd a) s4)) /\ + env (fst a) = Tfloat /\ + env (snd a) = Tfloat) \/ + f a (s1, s2, s3, s4) = (s1,s2, Regset.add (fst a) s3, Regset.add (snd a) s4) \/ + f a (s1, s2, s3, s4) = (s1,s2, Regset.add (snd a) s3, Regset.add (fst a) s4) +). +unfold f. +destruct (env (snd a)); destruct (env (fst a)); unfold rr1, rr2, rr3, rr4; simpl. +left. auto. +right. right. right. reflexivity. +right. right. left. reflexivity. +right. left. auto. + +destruct H0. +destruct H0. destruct H1. rewrite H0 in H. +generalize (IHl (SetRegReg.add a s1) s2 _ _ H). intro H3. +destruct H3. +left. assumption. +destruct (proj1 (add_iff _ _ _) H3). +left. intuition. rewrite <-H5. auto. rewrite <-H6. auto. +right. assumption. + +destruct H0. destruct H0. destruct H1. rewrite H0 in H. +apply (IHl s1 (SetRegReg.add a s2) _ _ H). + +destruct H0. rewrite H0 in H. apply (IHl s1 s2 _ _ H). +rewrite H0 in H. apply (IHl s1 s2 _ _ H). +Qed. + +Lemma in_partition_type_snd : forall e s env, +SetRegReg.In e (rr2 (regreg_edge_type_partition s env)) -> +env (fst e) = Tfloat /\ env (snd e) = Tfloat. + +Proof. +intros e s env H. +unfold regreg_edge_type_partition in H. +set (f := fun (e : SetRegReg.elt) (s : regregpartition) => + match env (fst e) with + | Tint => + match env (snd e) with + | Tint => + (SetRegReg.add e (rr1 s), rr2 s, + Regset.add (fst e) (Regset.add (snd e) (rr3 s)), + rr4 s) + | Tfloat => + (rr1 s, rr2 s, Regset.add (fst e) (rr3 s), + Regset.add (snd e) (rr4 s)) + end + | Tfloat => + match env (snd e) with + | Tint => + (rr1 s, rr2 s, Regset.add (snd e) (rr3 s), + Regset.add (fst e) (rr4 s)) + | Tfloat => + (rr1 s, SetRegReg.add e (rr2 s), rr3 s, + Regset.add (fst e) (Regset.add (snd e) (rr4 s))) + end + end). +unfold regregpartition in *. fold f in H. + +cut (forall e set1 set2 set3 set4, SetRegReg.In e (rr2 (SetRegReg.fold f s (set1, set2, set3, set4))) -> + (env (fst e) = Tfloat /\ env (snd e) = Tfloat) \/ SetRegReg.In e set2). + +intros. +generalize (H0 _ _ _ _ _ H). clear H H0. intro H. +destruct H. assumption. elim (SetRegReg.empty_1 H). + +(* cut property *) +intros. +generalize H0. clear H H0. intro H. +rewrite SetRegReg.fold_1 in H. +generalize H. clear H. generalize set1 set2 set3 set4. +induction (SetRegReg.elements s); intros s1 s2 s3 s4 H. +simpl in H. right. assumption. +simpl in H. + +assert ((f a (s1, s2, s3, s4) = (SetRegReg.add a s1, s2, Regset.add (fst a) (Regset.add (snd a) s3), s4) /\ + env (fst a) = Tint /\ + env (snd a) = Tint)\/ + (f a (s1, s2, s3, s4) = (s1, SetRegReg.add a s2, s3, Regset.add (fst a) (Regset.add (snd a) s4)) /\ + env (fst a) = Tfloat /\ + env (snd a) = Tfloat) \/ + f a (s1, s2, s3, s4) = (s1,s2, Regset.add (fst a) s3, Regset.add (snd a) s4) \/ + f a (s1, s2, s3, s4) = (s1,s2, Regset.add (snd a) s3, Regset.add (fst a) s4) +). +unfold f. +destruct (env (snd a)); destruct (env (fst a)); unfold rr1, rr2, rr3, rr4; simpl. +left. auto. +right. right. right. reflexivity. +right. right. left. reflexivity. +right. left. auto. + +destruct H0. +destruct H0. rewrite H0 in H. +apply (IHl (SetRegReg.add a s1) s2 _ _ H). +destruct H0. destruct H0. rewrite H0 in H. +generalize (IHl s1 (SetRegReg.add a s2) _ _ H). intro. +destruct H2. +left. assumption. +destruct (proj1 (add_iff _ _ _) H2). +left. intuition. rewrite <-H1. auto. rewrite <-H6. auto. +right. assumption. + +destruct H0. rewrite H0 in H. apply (IHl _ _ _ _ H). + +rewrite H0 in H. apply (IHl s1 s2 _ _ H). +Qed. + +Definition regmregpartition : Type := SetRegMreg.t*SetRegMreg.t*Regset.t*Regset.t*MRegset.t*MRegset.t. + +Definition rm1 := fun (p : regmregpartition) => fst (fst (fst (fst (fst p)))). +Definition rm2 := fun (p : regmregpartition) => snd (fst (fst (fst (fst p)))). +Definition rm3 := fun (p : regmregpartition) => snd (fst (fst (fst p))). +Definition rm4 := fun (p : regmregpartition) => snd (fst (fst p)). +Definition rm5 := fun (p : regmregpartition) => snd (fst p). +Definition rm6 := fun (p : regmregpartition) => snd p. + +Module Import SRMFacts := Facts SetRegMreg. + +Definition regmreg_edge_type_partition s env := +SetRegMreg.fold (fun e s => match env (fst e), mreg_type (snd e) with + | Tint, Tint => (SetRegMreg.add e (rm1 s), rm2 s, + Regset.add (fst e) (rm3 s), rm4 s, + MRegset.add (snd e) (rm5 s), rm6 s) + | Tfloat, Tfloat => (rm1 s, SetRegMreg.add e (rm2 s), + rm3 s, Regset.add (fst e) (rm4 s), + rm5 s, MRegset.add (snd e) (rm6 s)) + | Tint, Tfloat => (rm1 s, rm2 s, + Regset.add (fst e) (rm3 s), rm4 s, + rm5 s, MRegset.add (snd e) (rm6 s)) + |Tfloat, Tint => (rm1 s, rm2 s, + rm3 s, Regset.add (fst e) (rm4 s), + MRegset.add (snd e) (rm5 s), rm6 s) + end) +s +(SetRegMreg.empty, SetRegMreg.empty, Regset.empty, Regset.empty, MRegset.empty, MRegset.empty). + +Lemma in_mreg_partition_in_fst : forall e s env, +SetRegMreg.In e (rm1 (regmreg_edge_type_partition s env)) -> +SetRegMreg.In e s. + +Proof. +Admitted. +(* +intros e s env H. +unfold regmreg_edge_type_partition in H. +set (f := (fun (e : SetRegMreg.elt) (s : SetRegMreg.t * SetRegMreg.t) => + match env (fst e) with + | Tint => + match mreg_type (snd e) with + | Tint => (SetRegMreg.add e (fst s), snd s) + | Tfloat => s + end + | Tfloat => + match mreg_type (snd e) with + | Tint => s + | Tfloat => (fst s, SetRegMreg.add e (snd s)) + end + end)) in H. + +assert (forall e set1 set2, SetRegMreg.In e (fst (SetRegMreg.fold f s (set1, set2))) -> + SetRegMreg.In e s \/ SetRegMreg.In e set1 \/ SetRegMreg.In e set2). +clear H. +intros e' s1 s2 H. +rewrite SetRegMreg.fold_1 in H. +generalize H. generalize s1 s2. clear s1 s2 H. +generalize (SetRegMreg.elements_2). intro HH. +generalize (HH s). clear HH. intro HH. +induction (SetRegMreg.elements s). +simpl. right. left. assumption. +intros s1 s2 H. +simpl in H. +assert ((forall x : SetRegMreg.elt, + SetoidList.InA (fun x0 y : OrderedRegMreg.t => fst x0 = fst y /\ snd x0 = snd y) x l -> + SetRegMreg.In x s)). +intros. apply HH. right. assumption. +generalize (IHl H0). clear IHl H0. intro IHl. +assert (f a (s1, s2) = (SetRegMreg.add a s1, s2) \/ + f a (s1, s2) = (s1, SetRegMreg.add a s2) \/ + f a (s1, s2) = (s1,s2)). +unfold f. +destruct (mreg_type (snd a)); destruct (env (fst a)); simpl. +left. reflexivity. +right. right. reflexivity. +right. right. reflexivity. +right. left. reflexivity. +destruct H0. +rewrite H0 in H. + +generalize (IHl (SetRegMreg.add a s1) s2 H). +intro H1. destruct H1. +left. assumption. +destruct H1. + +destruct (proj1 (add_iff _ _ _) H1). +left. apply HH. left. intuition. +right. left. assumption. +right. right. assumption. +destruct H0. +rewrite H0 in H. +generalize (IHl s1 (SetRegMreg.add a s2) H). +intro H1. destruct H1. +left. assumption. +destruct H1. +right. left. assumption. +destruct (proj1 (add_iff _ _ _) H1). +left. apply HH. left. intuition. +right. right. assumption. +rewrite H0 in H. +generalize (IHl s1 s2 H). +intro H1. destruct H1. +left. assumption. +destruct H1. +right. left. assumption. +right. right. assumption. +generalize (H0 _ _ _ H). clear H0. intro H0. +destruct H0. +assumption. +destruct H0; elim (SetRegMreg.empty_1 H0). +Qed. +*) + +Lemma in_mreg_partition_in_snd : forall e s env, +SetRegMreg.In e (rm2 (regmreg_edge_type_partition s env)) -> +SetRegMreg.In e s. + +Proof. +Admitted. +(* +intros e s env H. +unfold regmreg_edge_type_partition in H. +set (f := (fun (e : SetRegMreg.elt) (s : SetRegMreg.t * SetRegMreg.t) => + match env (fst e) with + | Tint => + match mreg_type (snd e) with + | Tint => (SetRegMreg.add e (fst s), snd s) + | Tfloat => s + end + | Tfloat => + match mreg_type (snd e) with + | Tint => s + | Tfloat => (fst s, SetRegMreg.add e (snd s)) + end + end)) in H. + +assert (forall e set1 set2, SetRegMreg.In e (snd (SetRegMreg.fold f s (set1, set2))) -> + SetRegMreg.In e s \/ SetRegMreg.In e set1 \/ SetRegMreg.In e set2). +clear H. +intros e' s1 s2 H. +rewrite SetRegMreg.fold_1 in H. +generalize H. generalize s1 s2. clear s1 s2 H. +generalize (SetRegMreg.elements_2). intro HH. +generalize (HH s). clear HH. intro HH. +induction (SetRegMreg.elements s). +simpl. right. right. assumption. +intros s1 s2 H. +simpl in H. +assert ((forall x : SetRegMreg.elt, + SetoidList.InA (fun x0 y : OrderedRegMreg.t => fst x0 = fst y /\ snd x0 = snd y) x l -> + SetRegMreg.In x s)). +intros. apply HH. right. assumption. +generalize (IHl H0). clear IHl H0. intro IHl. +assert (f a (s1, s2) = (SetRegMreg.add a s1, s2) \/ + f a (s1, s2) = (s1, SetRegMreg.add a s2) \/ + f a (s1, s2) = (s1,s2)). +unfold f. +destruct (mreg_type (snd a)); destruct (env (fst a)); simpl. +left. reflexivity. +right. right. reflexivity. +right. right. reflexivity. +right. left. reflexivity. +destruct H0. +rewrite H0 in H. + +generalize (IHl (SetRegMreg.add a s1) s2 H). +intro H1. destruct H1. +left. assumption. +destruct H1. + +destruct (proj1 (add_iff _ _ _) H1). +left. apply HH. left. intuition. +right. left. assumption. +right. right. assumption. +destruct H0. +rewrite H0 in H. +generalize (IHl s1 (SetRegMreg.add a s2) H). +intro H1. destruct H1. +left. assumption. +destruct H1. +right. left. assumption. +destruct (proj1 (add_iff _ _ _) H1). +left. apply HH. left. intuition. +right. right. assumption. +rewrite H0 in H. +generalize (IHl s1 s2 H). +intro H1. destruct H1. +left. assumption. +destruct H1. +right. left. assumption. +right. right. assumption. +generalize (H0 _ _ _ H). clear H0. intro H0. +destruct H0. +assumption. +destruct H0; elim (SetRegMreg.empty_1 H0). +Qed. +*) + +Lemma in_mreg_partition_type_fst : forall e s env, +SetRegMreg.In e (rm1 (regmreg_edge_type_partition s env)) -> +env (fst e) = Tint /\ mreg_type (snd e) = Tint. + +Proof. +Admitted. +(* +intros e s env H. +unfold regmreg_edge_type_partition in H. +set (f := (fun (e : SetRegMreg.elt) (s : SetRegMreg.t * SetRegMreg.t) => + match env (fst e) with + | Tint => + match mreg_type (snd e) with + | Tint => (SetRegMreg.add e (fst s), snd s) + | Tfloat => s + end + | Tfloat => + match mreg_type (snd e) with + | Tint => s + | Tfloat => (fst s, SetRegMreg.add e (snd s)) + end + end)) in H. + +cut (forall e set1 set2, SetRegMreg.In e (fst (SetRegMreg.fold f s (set1, set2))) -> + (env (fst e) = Tint /\ mreg_type (snd e) = Tint) \/ SetRegMreg.In e set1). + +intros. +generalize (H0 _ _ _ H). clear H H0. intro H. +destruct H. assumption. elim (SetRegMreg.empty_1 H). + +(* cut property *) +intros. +generalize H0. clear H H0. intro H. +rewrite SetRegMreg.fold_1 in H. +generalize H. clear H. generalize set1 set2. +induction (SetRegMreg.elements s); intros s1 s2 H. +simpl in H. right. assumption. +simpl in H. + +assert ((f a (s1, s2) = (SetRegMreg.add a s1, s2) /\ + env (fst a) = Tint /\ + mreg_type (snd a) = Tint)\/ + (f a (s1, s2) = (s1, SetRegMreg.add a s2) /\ + env (fst a) = Tfloat /\ + mreg_type (snd a) = Tfloat) \/ + f a (s1, s2) = (s1,s2)). +unfold f. +destruct (mreg_type (snd a)); destruct (env (fst a)); simpl. +left. auto. +right. right. reflexivity. +right. right. reflexivity. +right. left. auto. + +destruct H0. +destruct H0. destruct H1. rewrite H0 in H. +generalize (IHl (SetRegMreg.add a s1) s2 H). intro H3. +destruct H3. +left. assumption. +destruct (proj1 (add_iff _ _ _) H3). +left. intuition. rewrite <-H5. auto. rewrite <-H6. auto. +right. assumption. + +destruct H0. destruct H0. destruct H1. rewrite H0 in H. +apply (IHl s1 (SetRegMreg.add a s2) H). + +rewrite H0 in H. +apply (IHl s1 s2 H). +Qed. +*) + +Lemma in_mreg_partition_type_snd : forall e s env, +SetRegMreg.In e (rm2 (regmreg_edge_type_partition s env)) -> +env (fst e) = Tfloat /\ mreg_type (snd e) = Tfloat. + +Proof. +Admitted. +(* +intros e s env H. +unfold regmreg_edge_type_partition in H. +set (f := (fun (e : SetRegMreg.elt) (s : SetRegMreg.t * SetRegMreg.t) => + match env (fst e) with + | Tint => + match mreg_type (snd e) with + | Tint => (SetRegMreg.add e (fst s), snd s) + | Tfloat => s + end + | Tfloat => + match mreg_type (snd e) with + | Tint => s + | Tfloat => (fst s, SetRegMreg.add e (snd s)) + end + end)) in H. + +cut (forall e set1 set2, SetRegMreg.In e (snd (SetRegMreg.fold f s (set1, set2))) -> + (env (fst e) = Tfloat /\ mreg_type (snd e) = Tfloat) \/ SetRegMreg.In e set2). + +intros. +generalize (H0 _ _ _ H). clear H H0. intro H. +destruct H. assumption. elim (SetRegMreg.empty_1 H). + +(* cut property *) +intros. +generalize H0. clear H H0. intro H. +rewrite SetRegMreg.fold_1 in H. +generalize H. clear H. generalize set1 set2. +induction (SetRegMreg.elements s); intros s1 s2 H. +simpl in H. right. assumption. +simpl in H. + +assert ((f a (s1, s2) = (SetRegMreg.add a s1, s2) /\ + env (fst a) = Tint /\ + mreg_type (snd a) = Tint)\/ + (f a (s1, s2) = (s1, SetRegMreg.add a s2) /\ + env (fst a) = Tfloat /\ + mreg_type (snd a) = Tfloat) \/ + f a (s1, s2) = (s1,s2)). +unfold f. +destruct (mreg_type (snd a)); destruct (env (fst a)); simpl. +left. auto. +right. right. reflexivity. +right. right. reflexivity. +right. left. auto. + +destruct H0. +destruct H0. destruct H1. rewrite H0 in H. +apply (IHl (SetRegMreg.add a s1) s2 H). + +destruct H0. destruct H0. destruct H1. rewrite H0 in H. +generalize (IHl s1 (SetRegMreg.add a s2) H). intro. +destruct H3. +left. auto. +destruct (proj1 (add_iff _ _ _) H3). +destruct H4. left. rewrite H4 in *. rewrite H5 in *. intuition. +right. auto. +rewrite H0 in H. apply (IHl s1 s2 H). +Qed. + +Definition Typed_interfgraphs g env := +let (int_regreg_interf, float_regreg_interf) := + regreg_edge_type_partition (interf_reg_reg g) env in +let (int_regmreg_interf, float_regmreg_interf) := + regmreg_edge_type_partition (interf_reg_mreg g) env in +let (int_regreg_pref, float_regreg_pref) := + regreg_edge_type_partition (pref_reg_reg g) env in +let (int_regmreg_pref, float_regmreg_pref) := + regmreg_edge_type_partition (pref_reg_mreg g) env in +(mkgraph int_regreg_interf int_regmreg_interf int_regreg_pref int_regmreg_pref, + mkgraph float_regreg_interf float_regmreg_interf float_regreg_pref float_regmreg_pref). +*) + + diff --git a/backend/WS.v b/backend/WS.v new file mode 100755 index 00000000..2f433272 --- /dev/null +++ b/backend/WS.v @@ -0,0 +1,231 @@ +Require Import FSets. +Require Import InterfGraphMapImp. +Require Import Conservative_criteria. +Require Import Edges. +Require Import MyRegisters. +Require Import Affinity_relation. + +Import Edge RegFacts Props OTFacts. + +(* Intersections of vertices sets of the worklists are empty *) +Lemma WS_empty_inter_1 : forall g palette WL, +WS_properties g palette WL -> +VertexSet.Empty (VertexSet.inter (get_spillWL WL) (get_freezeWL WL)). + +Proof. +intros g palette WL H. +unfold VertexSet.Empty. +intros a H0. +generalize (VertexSet.inter_1 H0);intro H1. +generalize (VertexSet.inter_2 H0);intro H2. +unfold WS_properties in H. +destruct H as [H H3];destruct H3 as [H3 H4];destruct H4 as [H4 H5]. +generalize (proj1 (H a) H1);intro H6. +destruct H6 as [H6 _]. +generalize (proj1 (H3 a) H2);intro H7. +destruct H7 as [H7 _]. +rewrite H6 in H7; inversion H7. +Qed. + +Lemma WS_empty_inter_2 : forall g palette WL, +WS_properties g palette WL -> +VertexSet.Empty (VertexSet.inter (get_spillWL WL) (get_simplifyWL WL)). + +Proof. +intros g palette WL H. +unfold VertexSet.Empty. +intros a H0. +generalize (VertexSet.inter_1 H0);intro H1. +generalize (VertexSet.inter_2 H0);intro H2. +unfold WS_properties in H. +destruct H as [H H3];destruct H3 as [H3 H4];destruct H4 as [H4 H5]. +generalize (proj1 (H a) H1);intro H6. +destruct H6 as [H6 _]. +generalize (proj1 (H4 a) H2);intro H7. +destruct H7 as [H7 _]. +rewrite H6 in H7; inversion H7. +Qed. + +Lemma WS_empty_inter_3 : forall g palette WL, +WS_properties g palette WL -> +VertexSet.Empty (VertexSet.inter (get_freezeWL WL) (get_simplifyWL WL)). + +Proof. +intros g palette WL H. +unfold VertexSet.Empty. +intros a H0. +generalize (VertexSet.inter_1 H0);intro H1. +generalize (VertexSet.inter_2 H0);intro H2. +unfold WS_properties in H. +destruct H as [H H3];destruct H3 as [H3 H4];destruct H4 as [H4 H5]. +generalize (proj1 (H3 a) H1);intro H6. +destruct H6 as [_ H6]. +destruct H6 as [H6 _]. +generalize (proj1 (H4 a) H2);intro H7. +destruct H7 as [_ H7]. +destruct H7 as [H7 _]. +rewrite H6 in H7; inversion H7. +Qed. + +(* A tactic for simplifying proofs of belonging of a vertex to a worklist *) +Ltac WS_apply H := generalize H;intro HWS_; +match goal with +| |- (VertexSet.In ?A _) => destruct HWS_ as [HWS_ HWS__]; + try (apply (proj2 (HWS_ A))); + destruct HWS__ as [HWS__ HWS___]; + try (apply (proj2 (HWS__ A))); + destruct HWS___ as [HWS___ HWS____]; + try (apply (proj2 (HWS___ A))); + clear HWS_ HWS__ HWS___ HWS____ +| |- (EdgeSet.In ?A _) => do 3 destruct HWS_ as [_ HWS_]; + apply (proj2 (HWS_ A)); + clear HWS_ +end. + +(* Lemmas for generalizing properties of a vertex belonging to a given worklist *) +Lemma In_spill_props : forall x g WL s a b c palette, +VertexSet.In x s -> +WL = (s,a,b,c) -> +WS_properties g palette WL -> +has_low_degree g palette x = false /\ In_graph x g /\ ~VertexSet.In x (precolored g). + +Proof. +intros x g WL s a b c palette H H0 H1. +unfold WS_properties in H1;rewrite H0 in H1. +destruct H1 as [H1 _]. +unfold get_spillWL in H1;simpl in H1. +apply (proj1 (H1 x) H). +Qed. + +Lemma In_freeze_props : forall x g WL s a b c palette, +VertexSet.In x s -> +WL = (a,s,b,c) -> +WS_properties g palette WL -> +has_low_degree g palette x = true /\ move_related g x = true /\ In_graph x g /\ ~VertexSet.In x (precolored g). + +Proof. +intros x g WL s a b c palette H H0 H1. +unfold WS_properties in H1;rewrite H0 in H1. +destruct H1 as [_ H1]. +destruct H1 as [H1 _]. +unfold get_freezeWL in H1;simpl in H1. +generalize (proj1 (H1 x) H);intro. +intuition. +apply move_related_in_graph;intuition. +Qed. + +Lemma In_simplify_props : forall x g WL s a b c palette, +VertexSet.In x s -> +WL = (a,b,s,c) -> +WS_properties g palette WL -> +has_low_degree g palette x = true /\ move_related g x = false /\ In_graph x g /\ ~VertexSet.In x (precolored g). + +Proof. +intros x g WL s a b c palette H H0 H1. +unfold WS_properties in H1;rewrite H0 in H1. +do 2 destruct H1 as [_ H1]. +destruct H1 as [H1 _]. +unfold get_spillWL in H1;simpl in H1. +apply (proj1 (H1 x) H). +Qed. + +Lemma In_move_props : forall e g WL s a b c palette, +EdgeSet.In e s -> +WL = (a,b,c,s) -> +WS_properties g palette WL -> +aff_edge e /\ In_graph_edge e g. + +Proof. +intros e g WL s a b c palette H H0 H1. +unfold WS_properties in H1;rewrite H0 in H1. +do 3 destruct H1 as [_ H1]. +unfold get_movesWL in H1;simpl in H1. +apply (proj1 (H1 e) H). +Qed. + +Lemma WS_props_equal : +forall g palette ws ws', +VertexSet.Equal (get_simplifyWL ws) (get_simplifyWL ws') -> +VertexSet.Equal (get_freezeWL ws) (get_freezeWL ws') -> +VertexSet.Equal (get_spillWL ws) (get_spillWL ws') -> +EdgeSet.Equal (get_movesWL ws) (get_movesWL ws') -> +WS_properties g palette ws -> +WS_properties g palette ws'. + +Proof. +unfold WS_properties;unfold get_spillWL;unfold get_freezeWL; +unfold get_simplifyWL;unfold get_movesWL;simpl;unfold VertexSet.Equal; +unfold EdgeSet.Equal;intros g palette ws ws' H H0 H1 H2 H3. +destruct ws as [ws d]; destruct ws as [ws c]; destruct ws as [a b]. +destruct ws' as [ws' p]; destruct ws' as [ws' o]; destruct ws' as [m n]. simpl in *. +generalize (VertexSet.eq_sym H);generalize (VertexSet.eq_sym H0); +generalize (VertexSet.eq_sym H1);generalize (EdgeSet.eq_sym H2); +clear H;clear H0;clear H1;clear H2;intros H2 H1 H0 H. + +destruct H3 as [Hsp H3];destruct H3 as [Hf H3]; +destruct H3 as [Hsi Hm]. +do 2 split. +intro H4;apply (proj1 (Hsp x) (proj1 (H1 x) H4)). +intro H4;apply (proj2 (H1 x) (proj2 (Hsp x) H4)). +split;intro H4. +apply (proj1 (Hf x) (proj1 (H0 x) H4)). +apply (proj2 (H0 x) (proj2 (Hf x) H4)). +split. +split;intro H4. +apply (proj1 (Hsi x) (proj1 (H x) H4)). +apply (proj2 (H x) (proj2 (Hsi x) H4)). +split;intro H4. +apply (proj1 (Hm e) (proj1 (H2 e) H4)). +apply (proj2 (H2 e) (proj2 (Hm e) H4)). +Qed. + +(* Definition of the nonprecolored vertices of a graph *) +Definition not_precolored g := VertexSet.diff (V g) (precolored g). + +(* The union of vertices worklists forms the nonprecolored vertices set of g *) +Lemma not_precolored_union_ws : forall g palette ws, +WS_properties g palette ws -> +VertexSet.Equal +(VertexSet.union (VertexSet.union (get_spillWL ws) (get_freezeWL ws)) (get_simplifyWL ws)) +(not_precolored g). + +Proof. +intros g palette ws HWS. +split. intro H. +unfold not_precolored. apply VertexSet.diff_3. +destruct (VertexSet.union_1 H). +destruct (VertexSet.union_1 H0). +apply (proj1(proj2 (In_spill_props _ _ _ _ _ _ _ _ H1 (refl_equal _) HWS))). +apply (proj2(proj2 (In_freeze_props _ _ _ _ _ _ _ _ H1 (refl_equal _) HWS))). +apply (proj2(proj2 (In_simplify_props _ _ _ _ _ _ _ _ H0 (refl_equal _) HWS))). +destruct (VertexSet.union_1 H). +destruct (VertexSet.union_1 H0). +apply (proj2(proj2 (In_spill_props _ _ _ _ _ _ _ _ H1 (refl_equal _) HWS))). +apply (proj2(proj2 (In_freeze_props _ _ _ _ _ _ _ _ H1 (refl_equal _) HWS))). +apply (proj2(proj2 (In_simplify_props _ _ _ _ _ _ _ _ H0 (refl_equal _) HWS))). +intro H. +unfold not_precolored in H. +generalize (VertexSet.diff_1 H). intro H0. +generalize (VertexSet.diff_2 H). clear H. intro H. +case_eq (has_low_degree g palette a); intro Hlow. +case_eq (move_related g a); intro Haff. +apply VertexSet.union_2. apply VertexSet.union_3. +WS_apply HWS. intuition. +apply VertexSet.union_3. +WS_apply HWS. intuition. +apply VertexSet.union_2. apply VertexSet.union_2. +WS_apply HWS. intuition. +Qed. + +(* The moves worklists is equal to the set of affinity edges *) +Lemma moves_AE : forall g palette ws, +WS_properties g palette ws -> +EdgeSet.Equal (AE g) (get_movesWL ws). + +Proof. +split; intros. +destruct ws. destruct p. destruct p. +simpl. WS_apply H. apply (proj1 (In_graph_aff_edge_in_AE _ _) H0). +generalize (In_move_props _ _ _ _ _ _ _ _ H0 (refl_equal _) H). intro H1. +apply (proj2 (In_graph_aff_edge_in_AE _ _) H1). +Qed. -- cgit