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