| 
73477
 | 
     1  | 
(*  Author:     Lars Hupel, TU München
  | 
| 
73326
 | 
     2  | 
*)
  | 
| 
 | 
     3  | 
  | 
| 
 | 
     4  | 
section \<open>Disjoint FSets\<close>
  | 
| 
 | 
     5  | 
  | 
| 
 | 
     6  | 
theory Disjoint_FSets
  | 
| 
 | 
     7  | 
  imports
  | 
| 
 | 
     8  | 
    "HOL-Library.Finite_Map"
  | 
| 
73477
 | 
     9  | 
    Disjoint_Sets
  | 
| 
73326
 | 
    10  | 
begin
  | 
| 
 | 
    11  | 
  | 
| 
 | 
    12  | 
context
  | 
| 
 | 
    13  | 
  includes fset.lifting
  | 
| 
 | 
    14  | 
begin
  | 
| 
 | 
    15  | 
  | 
| 
 | 
    16  | 
lift_definition fdisjnt :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" is disjnt .
  | 
| 
 | 
    17  | 
  | 
| 
 | 
    18  | 
lemma fdisjnt_alt_def: "fdisjnt M N \<longleftrightarrow> (M |\<inter>| N = {||})"
 | 
| 
 | 
    19  | 
by transfer (simp add: disjnt_def)
  | 
| 
 | 
    20  | 
  | 
| 
 | 
    21  | 
lemma fdisjnt_insert: "x |\<notin>| N \<Longrightarrow> fdisjnt M N \<Longrightarrow> fdisjnt (finsert x M) N"
  | 
| 
 | 
    22  | 
by transfer' (rule disjnt_insert)
  | 
| 
 | 
    23  | 
  | 
| 
 | 
    24  | 
lemma fdisjnt_subset_right: "N' |\<subseteq>| N \<Longrightarrow> fdisjnt M N \<Longrightarrow> fdisjnt M N'"
  | 
| 
 | 
    25  | 
unfolding fdisjnt_alt_def by auto
  | 
| 
 | 
    26  | 
  | 
| 
 | 
    27  | 
lemma fdisjnt_subset_left: "N' |\<subseteq>| N \<Longrightarrow> fdisjnt N M \<Longrightarrow> fdisjnt N' M"
  | 
| 
 | 
    28  | 
unfolding fdisjnt_alt_def by auto
  | 
| 
 | 
    29  | 
  | 
| 
 | 
    30  | 
lemma fdisjnt_union_right: "fdisjnt M A \<Longrightarrow> fdisjnt M B \<Longrightarrow> fdisjnt M (A |\<union>| B)"
  | 
| 
 | 
    31  | 
unfolding fdisjnt_alt_def by auto
  | 
| 
 | 
    32  | 
  | 
| 
 | 
    33  | 
lemma fdisjnt_union_left: "fdisjnt A M \<Longrightarrow> fdisjnt B M \<Longrightarrow> fdisjnt (A |\<union>| B) M"
  | 
| 
 | 
    34  | 
unfolding fdisjnt_alt_def by auto
  | 
| 
 | 
    35  | 
  | 
| 
 | 
    36  | 
lemma fdisjnt_swap: "fdisjnt M N \<Longrightarrow> fdisjnt N M"
  | 
| 
 | 
    37  | 
including fset.lifting by transfer' (auto simp: disjnt_def)
  | 
| 
 | 
    38  | 
  | 
| 
 | 
    39  | 
lemma distinct_append_fset:
  | 
| 
 | 
    40  | 
  assumes "distinct xs" "distinct ys" "fdisjnt (fset_of_list xs) (fset_of_list ys)"
  | 
| 
 | 
    41  | 
  shows "distinct (xs @ ys)"
  | 
| 
 | 
    42  | 
using assms
  | 
| 
 | 
    43  | 
by transfer' (simp add: disjnt_def)
  | 
| 
 | 
    44  | 
  | 
| 
 | 
    45  | 
lemma fdisjnt_contrI:
  | 
| 
 | 
    46  | 
  assumes "\<And>x. x |\<in>| M \<Longrightarrow> x |\<in>| N \<Longrightarrow> False"
  | 
| 
 | 
    47  | 
  shows "fdisjnt M N"
  | 
| 
 | 
    48  | 
using assms
  | 
| 
 | 
    49  | 
by transfer' (auto simp: disjnt_def)
  | 
| 
 | 
    50  | 
  | 
| 
 | 
    51  | 
lemma fdisjnt_Union_left: "fdisjnt (ffUnion S) T \<longleftrightarrow> fBall S (\<lambda>S. fdisjnt S T)"
  | 
| 
 | 
    52  | 
by transfer' (auto simp: disjnt_def)
  | 
| 
 | 
    53  | 
  | 
| 
 | 
    54  | 
lemma fdisjnt_Union_right: "fdisjnt T (ffUnion S) \<longleftrightarrow> fBall S (\<lambda>S. fdisjnt T S)"
  | 
| 
 | 
    55  | 
by transfer' (auto simp: disjnt_def)
  | 
| 
 | 
    56  | 
  | 
| 
 | 
    57  | 
lemma fdisjnt_ge_max: "fBall X (\<lambda>x. x > fMax Y) \<Longrightarrow> fdisjnt X Y"
  | 
| 
 | 
    58  | 
by transfer (auto intro: disjnt_ge_max)
  | 
| 
 | 
    59  | 
  | 
| 
 | 
    60  | 
end
  | 
| 
 | 
    61  | 
  | 
| 
 | 
    62  | 
(* FIXME should be provable without lifting *)
  | 
| 
 | 
    63  | 
lemma fmadd_disjnt: "fdisjnt (fmdom m) (fmdom n) \<Longrightarrow> m ++\<^sub>f n = n ++\<^sub>f m"
  | 
| 
 | 
    64  | 
unfolding fdisjnt_alt_def
  | 
| 
 | 
    65  | 
including fset.lifting fmap.lifting
  | 
| 
 | 
    66  | 
apply transfer
  | 
| 
 | 
    67  | 
apply (rule ext)
  | 
| 
 | 
    68  | 
apply (auto simp: map_add_def split: option.splits)
  | 
| 
 | 
    69  | 
done
  | 
| 
 | 
    70  | 
  | 
| 
 | 
    71  | 
end
  |