src/HOL/Library/FSet.thy
author blanchet
Mon, 21 Oct 2013 08:27:51 +0200
changeset 54178 d6dc359426b7
parent 54014 21dac9a60f0c
child 54258 adfc759263ab
permissions -rw-r--r--
more informative abort
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
     1
(*  Title:      HOL/Library/FSet.thy
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
     2
    Author:     Ondrej Kuncar, TU Muenchen
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
     3
    Author:     Cezary Kaliszyk and Christian Urban
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
     4
*)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
     5
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
     6
header {* Type of finite sets defined as a subtype of sets *}
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
     7
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
     8
theory FSet
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
     9
imports Main Conditionally_Complete_Lattices
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    10
begin
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    11
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    12
subsection {* Definition of the type *}
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    13
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    14
typedef 'a fset = "{A :: 'a set. finite A}"  morphisms fset Abs_fset
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    15
by auto
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    16
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    17
setup_lifting type_definition_fset
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    18
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    19
subsection {* Basic operations and type class instantiations *}
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    20
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    21
(* FIXME transfer and right_total vs. bi_total *)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    22
instantiation fset :: (finite) finite
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    23
begin
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    24
instance by default (transfer, simp)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    25
end
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    26
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    27
instantiation fset :: (type) "{bounded_lattice_bot, distrib_lattice, minus}"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    28
begin
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    29
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    30
interpretation lifting_syntax .
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    31
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    32
lift_definition bot_fset :: "'a fset" is "{}" parametric empty_transfer by simp 
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    33
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    34
lift_definition less_eq_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" is subset_eq parametric subset_transfer 
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    35
  by simp
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    36
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    37
definition less_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" where "xs < ys \<equiv> xs \<le> ys \<and> xs \<noteq> (ys::'a fset)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    38
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    39
lemma less_fset_transfer[transfer_rule]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    40
  assumes [transfer_rule]: "bi_unique A" 
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    41
  shows "((pcr_fset A) ===> (pcr_fset A) ===> op =) op \<subset> op <"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    42
  unfolding less_fset_def[abs_def] psubset_eq[abs_def] by transfer_prover
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    43
  
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    44
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    45
lift_definition sup_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" is union parametric union_transfer
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    46
  by simp
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    47
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    48
lift_definition inf_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" is inter parametric inter_transfer
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    49
  by simp
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    50
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    51
lift_definition minus_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" is minus parametric Diff_transfer
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    52
  by simp
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    53
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    54
instance
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    55
by default (transfer, auto)+
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    56
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    57
end
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    58
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    59
abbreviation fempty :: "'a fset" ("{||}") where "{||} \<equiv> bot"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    60
abbreviation fsubset_eq :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<subseteq>|" 50) where "xs |\<subseteq>| ys \<equiv> xs \<le> ys"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    61
abbreviation fsubset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<subset>|" 50) where "xs |\<subset>| ys \<equiv> xs < ys"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    62
abbreviation funion :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" (infixl "|\<union>|" 65) where "xs |\<union>| ys \<equiv> sup xs ys"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    63
abbreviation finter :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" (infixl "|\<inter>|" 65) where "xs |\<inter>| ys \<equiv> inf xs ys"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    64
abbreviation fminus :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" (infixl "|-|" 65) where "xs |-| ys \<equiv> minus xs ys"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    65
54014
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53969
diff changeset
    66
instantiation fset :: (equal) equal
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53969
diff changeset
    67
begin
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53969
diff changeset
    68
definition "HOL.equal A B \<longleftrightarrow> A |\<subseteq>| B \<and> B |\<subseteq>| A"
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53969
diff changeset
    69
instance by intro_classes (auto simp add: equal_fset_def)
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53969
diff changeset
    70
end 
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53969
diff changeset
    71
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    72
instantiation fset :: (type) conditionally_complete_lattice
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    73
begin
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    74
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    75
interpretation lifting_syntax .
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    76
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    77
lemma right_total_Inf_fset_transfer:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    78
  assumes [transfer_rule]: "bi_unique A" and [transfer_rule]: "right_total A"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    79
  shows "(set_rel (set_rel A) ===> set_rel A) 
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    80
    (\<lambda>S. if finite (Inter S \<inter> Collect (Domainp A)) then Inter S \<inter> Collect (Domainp A) else {}) 
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    81
      (\<lambda>S. if finite (Inf S) then Inf S else {})"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    82
    by transfer_prover
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    83
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    84
lemma Inf_fset_transfer:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    85
  assumes [transfer_rule]: "bi_unique A" and [transfer_rule]: "bi_total A"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    86
  shows "(set_rel (set_rel A) ===> set_rel A) (\<lambda>A. if finite (Inf A) then Inf A else {}) 
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    87
    (\<lambda>A. if finite (Inf A) then Inf A else {})"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    88
  by transfer_prover
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    89
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    90
lift_definition Inf_fset :: "'a fset set \<Rightarrow> 'a fset" is "\<lambda>A. if finite (Inf A) then Inf A else {}" 
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    91
parametric right_total_Inf_fset_transfer Inf_fset_transfer by simp
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    92
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    93
lemma Sup_fset_transfer:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    94
  assumes [transfer_rule]: "bi_unique A"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    95
  shows "(set_rel (set_rel A) ===> set_rel A) (\<lambda>A. if finite (Sup A) then Sup A else {})
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    96
  (\<lambda>A. if finite (Sup A) then Sup A else {})" by transfer_prover
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    97
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    98
lift_definition Sup_fset :: "'a fset set \<Rightarrow> 'a fset" is "\<lambda>A. if finite (Sup A) then Sup A else {}"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    99
parametric Sup_fset_transfer by simp
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   100
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   101
lemma finite_Sup: "\<exists>z. finite z \<and> (\<forall>a. a \<in> X \<longrightarrow> a \<le> z) \<Longrightarrow> finite (Sup X)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   102
by (auto intro: finite_subset)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   103
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   104
instance
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   105
proof 
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   106
  fix x z :: "'a fset"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   107
  fix X :: "'a fset set"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   108
  {
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   109
    assume "x \<in> X" "(\<And>a. a \<in> X \<Longrightarrow> z |\<subseteq>| a)" 
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   110
    then show "Inf X |\<subseteq>| x"  by transfer auto
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   111
  next
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   112
    assume "X \<noteq> {}" "(\<And>x. x \<in> X \<Longrightarrow> z |\<subseteq>| x)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   113
    then show "z |\<subseteq>| Inf X" by transfer (clarsimp, blast)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   114
  next
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   115
    assume "x \<in> X" "(\<And>a. a \<in> X \<Longrightarrow> a |\<subseteq>| z)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   116
    then show "x |\<subseteq>| Sup X" by transfer (auto intro!: finite_Sup)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   117
  next
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   118
    assume "X \<noteq> {}" "(\<And>x. x \<in> X \<Longrightarrow> x |\<subseteq>| z)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   119
    then show "Sup X |\<subseteq>| z" by transfer (clarsimp, blast)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   120
  }
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   121
qed
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   122
end
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   123
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   124
instantiation fset :: (finite) complete_lattice 
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   125
begin
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   126
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   127
lift_definition top_fset :: "'a fset" is UNIV parametric right_total_UNIV_transfer UNIV_transfer by simp
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   128
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   129
instance by default (transfer, auto)+
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   130
end
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   131
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   132
instantiation fset :: (finite) complete_boolean_algebra
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   133
begin
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   134
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   135
lift_definition uminus_fset :: "'a fset \<Rightarrow> 'a fset" is uminus 
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   136
  parametric right_total_Compl_transfer Compl_transfer by simp
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   137
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   138
instance by (default, simp_all only: INF_def SUP_def) (transfer, auto)+
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   139
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   140
end
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   141
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   142
abbreviation fUNIV :: "'a::finite fset" where "fUNIV \<equiv> top"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   143
abbreviation fuminus :: "'a::finite fset \<Rightarrow> 'a fset" ("|-| _" [81] 80) where "|-| x \<equiv> uminus x"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   144
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   145
subsection {* Other operations *}
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   146
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   147
lift_definition finsert :: "'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset" is insert parametric Lifting_Set.insert_transfer
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   148
  by simp
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   149
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   150
syntax
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   151
  "_insert_fset"     :: "args => 'a fset"  ("{|(_)|}")
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   152
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   153
translations
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   154
  "{|x, xs|}" == "CONST finsert x {|xs|}"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   155
  "{|x|}"     == "CONST finsert x {||}"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   156
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   157
lift_definition fmember :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<in>|" 50) is Set.member 
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   158
  parametric member_transfer by simp
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   159
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   160
abbreviation notin_fset :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<notin>|" 50) where "x |\<notin>| S \<equiv> \<not> (x |\<in>| S)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   161
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   162
context
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   163
begin
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   164
interpretation lifting_syntax .
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   165
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   166
lift_definition ffilter :: "('a \<Rightarrow> bool) \<Rightarrow> 'a fset \<Rightarrow> 'a fset" is Set.filter 
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   167
  parametric Lifting_Set.filter_transfer unfolding Set.filter_def by simp
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   168
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   169
lemma compose_rel_to_Domainp:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   170
  assumes "left_unique R"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   171
  assumes "(R ===> op=) P P'"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   172
  shows "(R OO Lifting.invariant P' OO R\<inverse>\<inverse>) x y \<longleftrightarrow> Domainp R x \<and> P x \<and> x = y"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   173
using assms unfolding OO_def conversep_iff Domainp_iff left_unique_def fun_rel_def invariant_def
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   174
by blast
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   175
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   176
lift_definition fPow :: "'a fset \<Rightarrow> 'a fset fset" is Pow parametric Pow_transfer 
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   177
by (subst compose_rel_to_Domainp [OF _ finite_transfer]) (auto intro: transfer_raw finite_subset 
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   178
  simp add: fset.pcr_cr_eq[symmetric] Domainp_set fset.domain_eq)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   179
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   180
lift_definition fcard :: "'a fset \<Rightarrow> nat" is card parametric card_transfer by simp
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   181
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   182
lift_definition fimage :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset" (infixr "|`|" 90) is image 
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   183
  parametric image_transfer by simp
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   184
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   185
lift_definition fthe_elem :: "'a fset \<Rightarrow> 'a" is the_elem ..
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   186
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   187
(* FIXME why is not invariant here unfolded ? *)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   188
lift_definition fbind :: "'a fset \<Rightarrow> ('a \<Rightarrow> 'b fset) \<Rightarrow> 'b fset" is Set.bind parametric bind_transfer
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   189
unfolding invariant_def Set.bind_def by clarsimp metis
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   190
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   191
lift_definition ffUnion :: "'a fset fset \<Rightarrow> 'a fset" is Union parametric Union_transfer
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   192
by (subst(asm) compose_rel_to_Domainp [OF _ finite_transfer])
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   193
  (auto intro: transfer_raw simp add: fset.pcr_cr_eq[symmetric] Domainp_set fset.domain_eq invariant_def)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   194
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   195
lift_definition fBall :: "'a fset \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" is Ball parametric Ball_transfer ..
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   196
lift_definition fBex :: "'a fset \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" is Bex parametric Bex_transfer ..
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   197
53963
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   198
lift_definition ffold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a fset \<Rightarrow> 'b" is Finite_Set.fold ..
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   199
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   200
subsection {* Transferred lemmas from Set.thy *}
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   201
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   202
lemmas fset_eqI = set_eqI[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   203
lemmas fset_eq_iff[no_atp] = set_eq_iff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   204
lemmas fBallI[intro!] = ballI[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   205
lemmas fbspec[dest?] = bspec[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   206
lemmas fBallE[elim] = ballE[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   207
lemmas fBexI[intro] = bexI[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   208
lemmas rev_fBexI[intro?] = rev_bexI[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   209
lemmas fBexCI = bexCI[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   210
lemmas fBexE[elim!] = bexE[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   211
lemmas fBall_triv[simp] = ball_triv[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   212
lemmas fBex_triv[simp] = bex_triv[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   213
lemmas fBex_triv_one_point1[simp] = bex_triv_one_point1[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   214
lemmas fBex_triv_one_point2[simp] = bex_triv_one_point2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   215
lemmas fBex_one_point1[simp] = bex_one_point1[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   216
lemmas fBex_one_point2[simp] = bex_one_point2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   217
lemmas fBall_one_point1[simp] = ball_one_point1[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   218
lemmas fBall_one_point2[simp] = ball_one_point2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   219
lemmas fBall_conj_distrib = ball_conj_distrib[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   220
lemmas fBex_disj_distrib = bex_disj_distrib[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   221
lemmas fBall_cong = ball_cong[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   222
lemmas fBex_cong = bex_cong[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   223
lemmas fsubsetI[intro!] = subsetI[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   224
lemmas fsubsetD[elim, intro?] = subsetD[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   225
lemmas rev_fsubsetD[no_atp,intro?] = rev_subsetD[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   226
lemmas fsubsetCE[no_atp,elim] = subsetCE[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   227
lemmas fsubset_eq[no_atp] = subset_eq[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   228
lemmas contra_fsubsetD[no_atp] = contra_subsetD[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   229
lemmas fsubset_refl = subset_refl[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   230
lemmas fsubset_trans = subset_trans[Transfer.transferred]
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   231
lemmas fset_rev_mp = set_rev_mp[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   232
lemmas fset_mp = set_mp[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   233
lemmas fsubset_not_fsubset_eq[code] = subset_not_subset_eq[Transfer.transferred]
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   234
lemmas eq_fmem_trans = eq_mem_trans[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   235
lemmas fsubset_antisym[intro!] = subset_antisym[Transfer.transferred]
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   236
lemmas fequalityD1 = equalityD1[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   237
lemmas fequalityD2 = equalityD2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   238
lemmas fequalityE = equalityE[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   239
lemmas fequalityCE[elim] = equalityCE[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   240
lemmas eqfset_imp_iff = eqset_imp_iff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   241
lemmas eqfelem_imp_iff = eqelem_imp_iff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   242
lemmas fempty_iff[simp] = empty_iff[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   243
lemmas fempty_fsubsetI[iff] = empty_subsetI[Transfer.transferred]
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   244
lemmas equalsffemptyI = equals0I[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   245
lemmas equalsffemptyD = equals0D[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   246
lemmas fBall_fempty[simp] = ball_empty[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   247
lemmas fBex_fempty[simp] = bex_empty[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   248
lemmas fPow_iff[iff] = Pow_iff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   249
lemmas fPowI = PowI[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   250
lemmas fPowD = PowD[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   251
lemmas fPow_bottom = Pow_bottom[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   252
lemmas fPow_top = Pow_top[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   253
lemmas fPow_not_fempty = Pow_not_empty[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   254
lemmas finter_iff[simp] = Int_iff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   255
lemmas finterI[intro!] = IntI[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   256
lemmas finterD1 = IntD1[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   257
lemmas finterD2 = IntD2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   258
lemmas finterE[elim!] = IntE[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   259
lemmas funion_iff[simp] = Un_iff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   260
lemmas funionI1[elim?] = UnI1[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   261
lemmas funionI2[elim?] = UnI2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   262
lemmas funionCI[intro!] = UnCI[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   263
lemmas funionE[elim!] = UnE[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   264
lemmas fminus_iff[simp] = Diff_iff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   265
lemmas fminusI[intro!] = DiffI[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   266
lemmas fminusD1 = DiffD1[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   267
lemmas fminusD2 = DiffD2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   268
lemmas fminusE[elim!] = DiffE[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   269
lemmas finsert_iff[simp] = insert_iff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   270
lemmas finsertI1 = insertI1[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   271
lemmas finsertI2 = insertI2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   272
lemmas finsertE[elim!] = insertE[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   273
lemmas finsertCI[intro!] = insertCI[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   274
lemmas fsubset_finsert_iff = subset_insert_iff[Transfer.transferred]
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   275
lemmas finsert_ident = insert_ident[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   276
lemmas fsingletonI[intro!,no_atp] = singletonI[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   277
lemmas fsingletonD[dest!,no_atp] = singletonD[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   278
lemmas fsingleton_iff = singleton_iff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   279
lemmas fsingleton_inject[dest!] = singleton_inject[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   280
lemmas fsingleton_finsert_inj_eq[iff,no_atp] = singleton_insert_inj_eq[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   281
lemmas fsingleton_finsert_inj_eq'[iff,no_atp] = singleton_insert_inj_eq'[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   282
lemmas fsubset_fsingletonD = subset_singletonD[Transfer.transferred]
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   283
lemmas fminus_single_finsert = diff_single_insert[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   284
lemmas fdoubleton_eq_iff = doubleton_eq_iff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   285
lemmas funion_fsingleton_iff = Un_singleton_iff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   286
lemmas fsingleton_funion_iff = singleton_Un_iff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   287
lemmas fimage_eqI[simp, intro] = image_eqI[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   288
lemmas fimageI = imageI[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   289
lemmas rev_fimage_eqI = rev_image_eqI[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   290
lemmas fimageE[elim!] = imageE[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   291
lemmas Compr_fimage_eq = Compr_image_eq[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   292
lemmas fimage_funion = image_Un[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   293
lemmas fimage_iff = image_iff[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   294
lemmas fimage_fsubset_iff[no_atp] = image_subset_iff[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   295
lemmas fimage_fsubsetI = image_subsetI[Transfer.transferred]
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   296
lemmas fimage_ident[simp] = image_ident[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   297
lemmas split_if_fmem1 = split_if_mem1[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   298
lemmas split_if_fmem2 = split_if_mem2[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   299
lemmas pfsubsetI[intro!,no_atp] = psubsetI[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   300
lemmas pfsubsetE[elim!,no_atp] = psubsetE[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   301
lemmas pfsubset_finsert_iff = psubset_insert_iff[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   302
lemmas pfsubset_eq = psubset_eq[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   303
lemmas pfsubset_imp_fsubset = psubset_imp_subset[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   304
lemmas pfsubset_trans = psubset_trans[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   305
lemmas pfsubsetD = psubsetD[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   306
lemmas pfsubset_fsubset_trans = psubset_subset_trans[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   307
lemmas fsubset_pfsubset_trans = subset_psubset_trans[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   308
lemmas pfsubset_imp_ex_fmem = psubset_imp_ex_mem[Transfer.transferred]
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   309
lemmas fimage_fPow_mono = image_Pow_mono[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   310
lemmas fimage_fPow_surj = image_Pow_surj[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   311
lemmas fsubset_finsertI = subset_insertI[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   312
lemmas fsubset_finsertI2 = subset_insertI2[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   313
lemmas fsubset_finsert = subset_insert[Transfer.transferred]
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   314
lemmas funion_upper1 = Un_upper1[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   315
lemmas funion_upper2 = Un_upper2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   316
lemmas funion_least = Un_least[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   317
lemmas finter_lower1 = Int_lower1[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   318
lemmas finter_lower2 = Int_lower2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   319
lemmas finter_greatest = Int_greatest[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   320
lemmas fminus_fsubset = Diff_subset[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   321
lemmas fminus_fsubset_conv = Diff_subset_conv[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   322
lemmas fsubset_fempty[simp] = subset_empty[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   323
lemmas not_pfsubset_fempty[iff] = not_psubset_empty[Transfer.transferred]
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   324
lemmas finsert_is_funion = insert_is_Un[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   325
lemmas finsert_not_fempty[simp] = insert_not_empty[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   326
lemmas fempty_not_finsert = empty_not_insert[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   327
lemmas finsert_absorb = insert_absorb[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   328
lemmas finsert_absorb2[simp] = insert_absorb2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   329
lemmas finsert_commute = insert_commute[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   330
lemmas finsert_fsubset[simp] = insert_subset[Transfer.transferred]
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   331
lemmas finsert_inter_finsert[simp] = insert_inter_insert[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   332
lemmas finsert_disjoint[simp,no_atp] = insert_disjoint[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   333
lemmas disjoint_finsert[simp,no_atp] = disjoint_insert[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   334
lemmas fimage_fempty[simp] = image_empty[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   335
lemmas fimage_finsert[simp] = image_insert[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   336
lemmas fimage_constant = image_constant[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   337
lemmas fimage_constant_conv = image_constant_conv[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   338
lemmas fimage_fimage = image_image[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   339
lemmas finsert_fimage[simp] = insert_image[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   340
lemmas fimage_is_fempty[iff] = image_is_empty[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   341
lemmas fempty_is_fimage[iff] = empty_is_image[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   342
lemmas fimage_cong = image_cong[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   343
lemmas fimage_finter_fsubset = image_Int_subset[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   344
lemmas fimage_fminus_fsubset = image_diff_subset[Transfer.transferred]
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   345
lemmas finter_absorb = Int_absorb[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   346
lemmas finter_left_absorb = Int_left_absorb[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   347
lemmas finter_commute = Int_commute[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   348
lemmas finter_left_commute = Int_left_commute[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   349
lemmas finter_assoc = Int_assoc[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   350
lemmas finter_ac = Int_ac[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   351
lemmas finter_absorb1 = Int_absorb1[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   352
lemmas finter_absorb2 = Int_absorb2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   353
lemmas finter_fempty_left = Int_empty_left[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   354
lemmas finter_fempty_right = Int_empty_right[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   355
lemmas disjoint_iff_fnot_equal = disjoint_iff_not_equal[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   356
lemmas finter_funion_distrib = Int_Un_distrib[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   357
lemmas finter_funion_distrib2 = Int_Un_distrib2[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   358
lemmas finter_fsubset_iff[no_atp, simp] = Int_subset_iff[Transfer.transferred]
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   359
lemmas funion_absorb = Un_absorb[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   360
lemmas funion_left_absorb = Un_left_absorb[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   361
lemmas funion_commute = Un_commute[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   362
lemmas funion_left_commute = Un_left_commute[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   363
lemmas funion_assoc = Un_assoc[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   364
lemmas funion_ac = Un_ac[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   365
lemmas funion_absorb1 = Un_absorb1[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   366
lemmas funion_absorb2 = Un_absorb2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   367
lemmas funion_fempty_left = Un_empty_left[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   368
lemmas funion_fempty_right = Un_empty_right[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   369
lemmas funion_finsert_left[simp] = Un_insert_left[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   370
lemmas funion_finsert_right[simp] = Un_insert_right[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   371
lemmas finter_finsert_left = Int_insert_left[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   372
lemmas finter_finsert_left_ifffempty[simp] = Int_insert_left_if0[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   373
lemmas finter_finsert_left_if1[simp] = Int_insert_left_if1[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   374
lemmas finter_finsert_right = Int_insert_right[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   375
lemmas finter_finsert_right_ifffempty[simp] = Int_insert_right_if0[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   376
lemmas finter_finsert_right_if1[simp] = Int_insert_right_if1[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   377
lemmas funion_finter_distrib = Un_Int_distrib[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   378
lemmas funion_finter_distrib2 = Un_Int_distrib2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   379
lemmas funion_finter_crazy = Un_Int_crazy[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   380
lemmas fsubset_funion_eq = subset_Un_eq[Transfer.transferred]
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   381
lemmas funion_fempty[iff] = Un_empty[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   382
lemmas funion_fsubset_iff[no_atp, simp] = Un_subset_iff[Transfer.transferred]
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   383
lemmas funion_fminus_finter = Un_Diff_Int[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   384
lemmas fminus_finter2 = Diff_Int2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   385
lemmas funion_finter_assoc_eq = Un_Int_assoc_eq[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   386
lemmas fBall_funion = ball_Un[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   387
lemmas fBex_funion = bex_Un[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   388
lemmas fminus_eq_fempty_iff[simp,no_atp] = Diff_eq_empty_iff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   389
lemmas fminus_cancel[simp] = Diff_cancel[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   390
lemmas fminus_idemp[simp] = Diff_idemp[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   391
lemmas fminus_triv = Diff_triv[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   392
lemmas fempty_fminus[simp] = empty_Diff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   393
lemmas fminus_fempty[simp] = Diff_empty[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   394
lemmas fminus_finsertffempty[simp,no_atp] = Diff_insert0[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   395
lemmas fminus_finsert = Diff_insert[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   396
lemmas fminus_finsert2 = Diff_insert2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   397
lemmas finsert_fminus_if = insert_Diff_if[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   398
lemmas finsert_fminus1[simp] = insert_Diff1[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   399
lemmas finsert_fminus_single[simp] = insert_Diff_single[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   400
lemmas finsert_fminus = insert_Diff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   401
lemmas fminus_finsert_absorb = Diff_insert_absorb[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   402
lemmas fminus_disjoint[simp] = Diff_disjoint[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   403
lemmas fminus_partition = Diff_partition[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   404
lemmas double_fminus = double_diff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   405
lemmas funion_fminus_cancel[simp] = Un_Diff_cancel[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   406
lemmas funion_fminus_cancel2[simp] = Un_Diff_cancel2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   407
lemmas fminus_funion = Diff_Un[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   408
lemmas fminus_finter = Diff_Int[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   409
lemmas funion_fminus = Un_Diff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   410
lemmas finter_fminus = Int_Diff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   411
lemmas fminus_finter_distrib = Diff_Int_distrib[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   412
lemmas fminus_finter_distrib2 = Diff_Int_distrib2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   413
lemmas fUNIV_bool[no_atp] = UNIV_bool[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   414
lemmas fPow_fempty[simp] = Pow_empty[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   415
lemmas fPow_finsert = Pow_insert[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   416
lemmas funion_fPow_fsubset = Un_Pow_subset[Transfer.transferred]
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   417
lemmas fPow_finter_eq[simp] = Pow_Int_eq[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   418
lemmas fset_eq_fsubset = set_eq_subset[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   419
lemmas fsubset_iff[no_atp] = subset_iff[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   420
lemmas fsubset_iff_pfsubset_eq = subset_iff_psubset_eq[Transfer.transferred]
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   421
lemmas all_not_fin_conv[simp] = all_not_in_conv[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   422
lemmas ex_fin_conv = ex_in_conv[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   423
lemmas fimage_mono = image_mono[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   424
lemmas fPow_mono = Pow_mono[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   425
lemmas finsert_mono = insert_mono[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   426
lemmas funion_mono = Un_mono[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   427
lemmas finter_mono = Int_mono[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   428
lemmas fminus_mono = Diff_mono[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   429
lemmas fin_mono = in_mono[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   430
lemmas fthe_felem_eq[simp] = the_elem_eq[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   431
lemmas fLeast_mono = Least_mono[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   432
lemmas fbind_fbind = bind_bind[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   433
lemmas fempty_fbind[simp] = empty_bind[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   434
lemmas nonfempty_fbind_const = nonempty_bind_const[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   435
lemmas fbind_const = bind_const[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   436
lemmas ffmember_filter[simp] = member_filter[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   437
lemmas fequalityI = equalityI[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   438
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   439
subsection {* Additional lemmas*}
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   440
53969
7ed81754b069 proper document markup;
wenzelm
parents: 53964
diff changeset
   441
subsubsection {* @{text fsingleton} *}
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   442
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   443
lemmas fsingletonE = fsingletonD [elim_format]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   444
53969
7ed81754b069 proper document markup;
wenzelm
parents: 53964
diff changeset
   445
subsubsection {* @{text femepty} *}
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   446
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   447
lemma fempty_ffilter[simp]: "ffilter (\<lambda>_. False) A = {||}"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   448
by transfer auto
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   449
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   450
(* FIXME, transferred doesn't work here *)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   451
lemma femptyE [elim!]: "a |\<in>| {||} \<Longrightarrow> P"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   452
  by simp
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   453
53969
7ed81754b069 proper document markup;
wenzelm
parents: 53964
diff changeset
   454
subsubsection {* @{text fset} *}
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   455
53963
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   456
lemmas fset_simps[simp] = bot_fset.rep_eq finsert.rep_eq
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   457
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   458
lemma finite_fset [simp]: 
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   459
  shows "finite (fset S)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   460
  by transfer simp
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   461
53963
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   462
lemmas fset_cong = fset_inject
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   463
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   464
lemma filter_fset [simp]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   465
  shows "fset (ffilter P xs) = Collect P \<inter> fset xs"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   466
  by transfer auto
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   467
53963
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   468
lemma notin_fset: "x |\<notin>| S \<longleftrightarrow> x \<notin> fset S" by (simp add: fmember.rep_eq)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   469
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   470
lemmas inter_fset[simp] = inf_fset.rep_eq
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   471
53963
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   472
lemmas union_fset[simp] = sup_fset.rep_eq
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   473
53963
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   474
lemmas minus_fset[simp] = minus_fset.rep_eq
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   475
53969
7ed81754b069 proper document markup;
wenzelm
parents: 53964
diff changeset
   476
subsubsection {* @{text filter_fset} *}
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   477
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   478
lemma subset_ffilter: 
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   479
  "ffilter P A |\<subseteq>| ffilter Q A = (\<forall> x. x |\<in>| A \<longrightarrow> P x \<longrightarrow> Q x)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   480
  by transfer auto
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   481
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   482
lemma eq_ffilter: 
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   483
  "(ffilter P A = ffilter Q A) = (\<forall>x. x |\<in>| A \<longrightarrow> P x = Q x)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   484
  by transfer auto
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   485
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   486
lemma pfsubset_ffilter:
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   487
  "(\<And>x. x |\<in>| A \<Longrightarrow> P x \<Longrightarrow> Q x) \<Longrightarrow> (x |\<in>| A & \<not> P x & Q x) \<Longrightarrow> 
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   488
    ffilter P A |\<subset>| ffilter Q A"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   489
  unfolding less_fset_def by (auto simp add: subset_ffilter eq_ffilter)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   490
53969
7ed81754b069 proper document markup;
wenzelm
parents: 53964
diff changeset
   491
subsubsection {* @{text finsert} *}
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   492
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   493
(* FIXME, transferred doesn't work here *)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   494
lemma set_finsert:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   495
  assumes "x |\<in>| A"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   496
  obtains B where "A = finsert x B" and "x |\<notin>| B"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   497
using assms by transfer (metis Set.set_insert finite_insert)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   498
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   499
lemma mk_disjoint_finsert: "a |\<in>| A \<Longrightarrow> \<exists>B. A = finsert a B \<and> a |\<notin>| B"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   500
  by (rule_tac x = "A |-| {|a|}" in exI, blast)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   501
53969
7ed81754b069 proper document markup;
wenzelm
parents: 53964
diff changeset
   502
subsubsection {* @{text fimage} *}
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   503
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   504
lemma subset_fimage_iff: "(B |\<subseteq>| f|`|A) = (\<exists> AA. AA |\<subseteq>| A \<and> B = f|`|AA)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   505
by transfer (metis mem_Collect_eq rev_finite_subset subset_image_iff)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   506
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   507
subsubsection {* bounded quantification *}
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   508
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   509
lemma bex_simps [simp, no_atp]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   510
  "\<And>A P Q. fBex A (\<lambda>x. P x \<and> Q) = (fBex A P \<and> Q)" 
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   511
  "\<And>A P Q. fBex A (\<lambda>x. P \<and> Q x) = (P \<and> fBex A Q)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   512
  "\<And>P. fBex {||} P = False" 
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   513
  "\<And>a B P. fBex (finsert a B) P = (P a \<or> fBex B P)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   514
  "\<And>A P f. fBex (f |`| A) P = fBex A (\<lambda>x. P (f x))"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   515
  "\<And>A P. (\<not> fBex A P) = fBall A (\<lambda>x. \<not> P x)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   516
by auto
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   517
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   518
lemma ball_simps [simp, no_atp]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   519
  "\<And>A P Q. fBall A (\<lambda>x. P x \<or> Q) = (fBall A P \<or> Q)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   520
  "\<And>A P Q. fBall A (\<lambda>x. P \<or> Q x) = (P \<or> fBall A Q)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   521
  "\<And>A P Q. fBall A (\<lambda>x. P \<longrightarrow> Q x) = (P \<longrightarrow> fBall A Q)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   522
  "\<And>A P Q. fBall A (\<lambda>x. P x \<longrightarrow> Q) = (fBex A P \<longrightarrow> Q)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   523
  "\<And>P. fBall {||} P = True"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   524
  "\<And>a B P. fBall (finsert a B) P = (P a \<and> fBall B P)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   525
  "\<And>A P f. fBall (f |`| A) P = fBall A (\<lambda>x. P (f x))"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   526
  "\<And>A P. (\<not> fBall A P) = fBex A (\<lambda>x. \<not> P x)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   527
by auto
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   528
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   529
lemma atomize_fBall:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   530
    "(\<And>x. x |\<in>| A ==> P x) == Trueprop (fBall A (\<lambda>x. P x))"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   531
apply (simp only: atomize_all atomize_imp)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   532
apply (rule equal_intr_rule)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   533
by (transfer, simp)+
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   534
53963
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   535
end
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   536
53969
7ed81754b069 proper document markup;
wenzelm
parents: 53964
diff changeset
   537
subsubsection {* @{text fcard} *}
53963
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   538
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   539
(* FIXME: improve transferred to handle bounded meta quantification *)
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   540
53963
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   541
lemma fcard_fempty:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   542
  "fcard {||} = 0"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   543
  by transfer (rule card_empty)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   544
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   545
lemma fcard_finsert_disjoint:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   546
  "x |\<notin>| A \<Longrightarrow> fcard (finsert x A) = Suc (fcard A)"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   547
  by transfer (rule card_insert_disjoint)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   548
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   549
lemma fcard_finsert_if:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   550
  "fcard (finsert x A) = (if x |\<in>| A then fcard A else Suc (fcard A))"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   551
  by transfer (rule card_insert_if)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   552
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   553
lemma card_0_eq [simp, no_atp]:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   554
  "fcard A = 0 \<longleftrightarrow> A = {||}"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   555
  by transfer (rule card_0_eq)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   556
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   557
lemma fcard_Suc_fminus1:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   558
  "x |\<in>| A \<Longrightarrow> Suc (fcard (A |-| {|x|})) = fcard A"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   559
  by transfer (rule card_Suc_Diff1)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   560
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   561
lemma fcard_fminus_fsingleton:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   562
  "x |\<in>| A \<Longrightarrow> fcard (A |-| {|x|}) = fcard A - 1"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   563
  by transfer (rule card_Diff_singleton)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   564
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   565
lemma fcard_fminus_fsingleton_if:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   566
  "fcard (A |-| {|x|}) = (if x |\<in>| A then fcard A - 1 else fcard A)"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   567
  by transfer (rule card_Diff_singleton_if)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   568
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   569
lemma fcard_fminus_finsert[simp]:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   570
  assumes "a |\<in>| A" and "a |\<notin>| B"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   571
  shows "fcard (A |-| finsert a B) = fcard (A |-| B) - 1"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   572
using assms by transfer (rule card_Diff_insert)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   573
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   574
lemma fcard_finsert: "fcard (finsert x A) = Suc (fcard (A |-| {|x|}))"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   575
by transfer (rule card_insert)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   576
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   577
lemma fcard_finsert_le: "fcard A \<le> fcard (finsert x A)"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   578
by transfer (rule card_insert_le)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   579
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   580
lemma fcard_mono:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   581
  "A |\<subseteq>| B \<Longrightarrow> fcard A \<le> fcard B"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   582
by transfer (rule card_mono)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   583
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   584
lemma fcard_seteq: "A |\<subseteq>| B \<Longrightarrow> fcard B \<le> fcard A \<Longrightarrow> A = B"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   585
by transfer (rule card_seteq)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   586
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   587
lemma pfsubset_fcard_mono: "A |\<subset>| B \<Longrightarrow> fcard A < fcard B"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   588
by transfer (rule psubset_card_mono)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   589
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   590
lemma fcard_funion_finter: 
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   591
  "fcard A + fcard B = fcard (A |\<union>| B) + fcard (A |\<inter>| B)"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   592
by transfer (rule card_Un_Int)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   593
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   594
lemma fcard_funion_disjoint:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   595
  "A |\<inter>| B = {||} \<Longrightarrow> fcard (A |\<union>| B) = fcard A + fcard B"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   596
by transfer (rule card_Un_disjoint)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   597
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   598
lemma fcard_funion_fsubset:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   599
  "B |\<subseteq>| A \<Longrightarrow> fcard (A |-| B) = fcard A - fcard B"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   600
by transfer (rule card_Diff_subset)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   601
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   602
lemma diff_fcard_le_fcard_fminus:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   603
  "fcard A - fcard B \<le> fcard(A |-| B)"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   604
by transfer (rule diff_card_le_card_Diff)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   605
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   606
lemma fcard_fminus1_less: "x |\<in>| A \<Longrightarrow> fcard (A |-| {|x|}) < fcard A"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   607
by transfer (rule card_Diff1_less)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   608
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   609
lemma fcard_fminus2_less:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   610
  "x |\<in>| A \<Longrightarrow> y |\<in>| A \<Longrightarrow> fcard (A |-| {|x|} |-| {|y|}) < fcard A"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   611
by transfer (rule card_Diff2_less)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   612
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   613
lemma fcard_fminus1_le: "fcard (A |-| {|x|}) \<le> fcard A"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   614
by transfer (rule card_Diff1_le)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   615
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   616
lemma fcard_pfsubset: "A |\<subseteq>| B \<Longrightarrow> fcard A < fcard B \<Longrightarrow> A < B"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   617
by transfer (rule card_psubset)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   618
53969
7ed81754b069 proper document markup;
wenzelm
parents: 53964
diff changeset
   619
subsubsection {* @{text ffold} *}
53963
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   620
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   621
(* FIXME: improve transferred to handle bounded meta quantification *)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   622
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   623
context comp_fun_commute
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   624
begin
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   625
  lemmas ffold_empty[simp] = fold_empty[Transfer.transferred]
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   626
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   627
  lemma ffold_finsert [simp]:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   628
    assumes "x |\<notin>| A"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   629
    shows "ffold f z (finsert x A) = f x (ffold f z A)"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   630
    using assms by (transfer fixing: f) (rule fold_insert)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   631
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   632
  lemma ffold_fun_left_comm:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   633
    "f x (ffold f z A) = ffold f (f x z) A"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   634
    by (transfer fixing: f) (rule fold_fun_left_comm)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   635
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   636
  lemma ffold_finsert2:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   637
    "x |\<notin>| A \<Longrightarrow> ffold f z (finsert x A)  = ffold f (f x z) A"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   638
    by (transfer fixing: f) (rule fold_insert2)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   639
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   640
  lemma ffold_rec:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   641
    assumes "x |\<in>| A"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   642
    shows "ffold f z A = f x (ffold f z (A |-| {|x|}))"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   643
    using assms by (transfer fixing: f) (rule fold_rec)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   644
  
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   645
  lemma ffold_finsert_fremove:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   646
    "ffold f z (finsert x A) = f x (ffold f z (A |-| {|x|}))"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   647
     by (transfer fixing: f) (rule fold_insert_remove)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   648
end
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   649
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   650
lemma ffold_fimage:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   651
  assumes "inj_on g (fset A)"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   652
  shows "ffold f z (g |`| A) = ffold (f \<circ> g) z A"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   653
using assms by transfer' (rule fold_image)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   654
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   655
lemma ffold_cong:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   656
  assumes "comp_fun_commute f" "comp_fun_commute g"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   657
  "\<And>x. x |\<in>| A \<Longrightarrow> f x = g x"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   658
    and "s = t" and "A = B"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   659
  shows "ffold f s A = ffold g t B"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   660
using assms by transfer (metis Finite_Set.fold_cong)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   661
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   662
context comp_fun_idem
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   663
begin
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   664
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   665
  lemma ffold_finsert_idem:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   666
    "ffold f z (finsert x A)  = f x (ffold f z A)"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   667
    by (transfer fixing: f) (rule fold_insert_idem)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   668
  
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   669
  declare ffold_finsert [simp del] ffold_finsert_idem [simp]
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   670
  
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   671
  lemma ffold_finsert_idem2:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   672
    "ffold f z (finsert x A) = ffold f (f x z) A"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   673
    by (transfer fixing: f) (rule fold_insert_idem2)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   674
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   675
end
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   676
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   677
subsection {* Choice in fsets *}
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   678
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   679
lemma fset_choice: 
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   680
  assumes "\<forall>x. x |\<in>| A \<longrightarrow> (\<exists>y. P x y)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   681
  shows "\<exists>f. \<forall>x. x |\<in>| A \<longrightarrow> P x (f x)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   682
  using assms by transfer metis
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   683
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   684
subsection {* Induction and Cases rules for fsets *}
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   685
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   686
lemma fset_exhaust [case_names empty insert, cases type: fset]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   687
  assumes fempty_case: "S = {||} \<Longrightarrow> P" 
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   688
  and     finsert_case: "\<And>x S'. S = finsert x S' \<Longrightarrow> P"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   689
  shows "P"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   690
  using assms by transfer blast
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   691
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   692
lemma fset_induct [case_names empty insert]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   693
  assumes fempty_case: "P {||}"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   694
  and     finsert_case: "\<And>x S. P S \<Longrightarrow> P (finsert x S)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   695
  shows "P S"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   696
proof -
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   697
  (* FIXME transfer and right_total vs. bi_total *)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   698
  note Domainp_forall_transfer[transfer_rule]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   699
  show ?thesis
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   700
  using assms by transfer (auto intro: finite_induct)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   701
qed
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   702
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   703
lemma fset_induct_stronger [case_names empty insert, induct type: fset]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   704
  assumes empty_fset_case: "P {||}"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   705
  and     insert_fset_case: "\<And>x S. \<lbrakk>x |\<notin>| S; P S\<rbrakk> \<Longrightarrow> P (finsert x S)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   706
  shows "P S"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   707
proof -
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   708
  (* FIXME transfer and right_total vs. bi_total *)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   709
  note Domainp_forall_transfer[transfer_rule]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   710
  show ?thesis
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   711
  using assms by transfer (auto intro: finite_induct)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   712
qed
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   713
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   714
lemma fset_card_induct:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   715
  assumes empty_fset_case: "P {||}"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   716
  and     card_fset_Suc_case: "\<And>S T. Suc (fcard S) = (fcard T) \<Longrightarrow> P S \<Longrightarrow> P T"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   717
  shows "P S"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   718
proof (induct S)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   719
  case empty
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   720
  show "P {||}" by (rule empty_fset_case)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   721
next
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   722
  case (insert x S)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   723
  have h: "P S" by fact
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   724
  have "x |\<notin>| S" by fact
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   725
  then have "Suc (fcard S) = fcard (finsert x S)" 
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   726
    by transfer auto
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   727
  then show "P (finsert x S)" 
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   728
    using h card_fset_Suc_case by simp
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   729
qed
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   730
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   731
lemma fset_strong_cases:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   732
  obtains "xs = {||}"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   733
    | ys x where "x |\<notin>| ys" and "xs = finsert x ys"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   734
by transfer blast
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   735
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   736
lemma fset_induct2:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   737
  "P {||} {||} \<Longrightarrow>
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   738
  (\<And>x xs. x |\<notin>| xs \<Longrightarrow> P (finsert x xs) {||}) \<Longrightarrow>
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   739
  (\<And>y ys. y |\<notin>| ys \<Longrightarrow> P {||} (finsert y ys)) \<Longrightarrow>
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   740
  (\<And>x xs y ys. \<lbrakk>P xs ys; x |\<notin>| xs; y |\<notin>| ys\<rbrakk> \<Longrightarrow> P (finsert x xs) (finsert y ys)) \<Longrightarrow>
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   741
  P xsa ysa"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   742
  apply (induct xsa arbitrary: ysa)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   743
  apply (induct_tac x rule: fset_induct_stronger)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   744
  apply simp_all
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   745
  apply (induct_tac xa rule: fset_induct_stronger)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   746
  apply simp_all
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   747
  done
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   748
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   749
subsection {* Setup for Lifting/Transfer *}
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   750
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   751
subsubsection {* Relator and predicator properties *}
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   752
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   753
lift_definition fset_rel :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a fset \<Rightarrow> 'b fset \<Rightarrow> bool" is set_rel
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   754
parametric set_rel_transfer ..
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   755
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   756
lemma fset_rel_alt_def: "fset_rel R = (\<lambda>A B. (\<forall>x.\<exists>y. x|\<in>|A \<longrightarrow> y|\<in>|B \<and> R x y) 
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   757
  \<and> (\<forall>y. \<exists>x. y|\<in>|B \<longrightarrow> x|\<in>|A \<and> R x y))"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   758
apply (rule ext)+
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   759
apply transfer'
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   760
apply (subst set_rel_def[unfolded fun_eq_iff]) 
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   761
by blast
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   762
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   763
lemma fset_rel_conversep: "fset_rel (conversep R) = conversep (fset_rel R)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   764
  unfolding fset_rel_alt_def by auto
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   765
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   766
lemmas fset_rel_eq [relator_eq] = set_rel_eq[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   767
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   768
lemma fset_rel_mono[relator_mono]: "A \<le> B \<Longrightarrow> fset_rel A \<le> fset_rel B"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   769
unfolding fset_rel_alt_def by blast
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   770
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   771
lemma finite_set_rel:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   772
  assumes fin: "finite X" "finite Z"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   773
  assumes R_S: "set_rel (R OO S) X Z"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   774
  shows "\<exists>Y. finite Y \<and> set_rel R X Y \<and> set_rel S Y Z"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   775
proof -
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   776
  obtain f where f: "\<forall>x\<in>X. R x (f x) \<and> (\<exists>z\<in>Z. S (f x) z)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   777
  apply atomize_elim
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   778
  apply (subst bchoice_iff[symmetric])
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   779
  using R_S[unfolded set_rel_def OO_def] by blast
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   780
  
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   781
  obtain g where g: "\<forall>z\<in>Z. S (g z) z \<and> (\<exists>x\<in>X. R  x (g z))"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   782
  apply atomize_elim
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   783
  apply (subst bchoice_iff[symmetric])
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   784
  using R_S[unfolded set_rel_def OO_def] by blast
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   785
  
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   786
  let ?Y = "f ` X \<union> g ` Z"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   787
  have "finite ?Y" by (simp add: fin)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   788
  moreover have "set_rel R X ?Y"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   789
    unfolding set_rel_def
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   790
    using f g by clarsimp blast
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   791
  moreover have "set_rel S ?Y Z"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   792
    unfolding set_rel_def
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   793
    using f g by clarsimp blast
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   794
  ultimately show ?thesis by metis
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   795
qed
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   796
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   797
lemma fset_rel_OO[relator_distr]: "fset_rel R OO fset_rel S = fset_rel (R OO S)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   798
apply (rule ext)+
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   799
by transfer (auto intro: finite_set_rel set_rel_OO[unfolded fun_eq_iff, rule_format, THEN iffD1])
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   800
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   801
lemma Domainp_fset[relator_domain]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   802
  assumes "Domainp T = P"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   803
  shows "Domainp (fset_rel T) = (\<lambda>A. fBall A P)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   804
proof -
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   805
  from assms obtain f where f: "\<forall>x\<in>Collect P. T x (f x)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   806
    unfolding Domainp_iff[abs_def]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   807
    apply atomize_elim
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   808
    by (subst bchoice_iff[symmetric]) auto
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   809
  from assms f show ?thesis
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   810
    unfolding fun_eq_iff fset_rel_alt_def Domainp_iff
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   811
    apply clarify
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   812
    apply (rule iffI)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   813
      apply blast
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   814
    by (rename_tac A, rule_tac x="f |`| A" in exI, blast)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   815
qed
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   816
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   817
lemmas reflp_fset_rel[reflexivity_rule] = reflp_set_rel[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   818
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   819
lemma right_total_fset_rel[transfer_rule]: "right_total A \<Longrightarrow> right_total (fset_rel A)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   820
unfolding right_total_def 
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   821
apply transfer
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   822
apply (subst(asm) choice_iff)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   823
apply clarsimp
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   824
apply (rename_tac A f y, rule_tac x = "f ` y" in exI)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   825
by (auto simp add: set_rel_def)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   826
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   827
lemma left_total_fset_rel[reflexivity_rule]: "left_total A \<Longrightarrow> left_total (fset_rel A)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   828
unfolding left_total_def 
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   829
apply transfer
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   830
apply (subst(asm) choice_iff)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   831
apply clarsimp
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   832
apply (rename_tac A f y, rule_tac x = "f ` y" in exI)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   833
by (auto simp add: set_rel_def)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   834
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   835
lemmas right_unique_fset_rel[transfer_rule] = right_unique_set_rel[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   836
lemmas left_unique_fset_rel[reflexivity_rule] = left_unique_set_rel[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   837
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   838
thm right_unique_fset_rel left_unique_fset_rel
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   839
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   840
lemma bi_unique_fset_rel[transfer_rule]: "bi_unique A \<Longrightarrow> bi_unique (fset_rel A)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   841
by (auto intro: right_unique_fset_rel left_unique_fset_rel iff: bi_unique_iff)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   842
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   843
lemma bi_total_fset_rel[transfer_rule]: "bi_total A \<Longrightarrow> bi_total (fset_rel A)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   844
by (auto intro: right_total_fset_rel left_total_fset_rel iff: bi_total_iff)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   845
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   846
lemmas fset_invariant_commute [invariant_commute] = set_invariant_commute[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   847
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   848
subsubsection {* Quotient theorem for the Lifting package *}
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   849
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   850
lemma Quotient_fset_map[quot_map]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   851
  assumes "Quotient R Abs Rep T"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   852
  shows "Quotient (fset_rel R) (fimage Abs) (fimage Rep) (fset_rel T)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   853
  using assms unfolding Quotient_alt_def4
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   854
  by (simp add: fset_rel_OO[symmetric] fset_rel_conversep) (simp add: fset_rel_alt_def, blast)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   855
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   856
subsubsection {* Transfer rules for the Transfer package *}
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   857
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   858
text {* Unconditional transfer rules *}
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   859
53963
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   860
context
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   861
begin
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   862
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   863
interpretation lifting_syntax .
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   864
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   865
lemmas fempty_transfer [transfer_rule] = empty_transfer[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   866
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   867
lemma finsert_transfer [transfer_rule]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   868
  "(A ===> fset_rel A ===> fset_rel A) finsert finsert"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   869
  unfolding fun_rel_def fset_rel_alt_def by blast
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   870
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   871
lemma funion_transfer [transfer_rule]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   872
  "(fset_rel A ===> fset_rel A ===> fset_rel A) funion funion"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   873
  unfolding fun_rel_def fset_rel_alt_def by blast
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   874
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   875
lemma ffUnion_transfer [transfer_rule]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   876
  "(fset_rel (fset_rel A) ===> fset_rel A) ffUnion ffUnion"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   877
  unfolding fun_rel_def fset_rel_alt_def by transfer (simp, fast)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   878
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   879
lemma fimage_transfer [transfer_rule]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   880
  "((A ===> B) ===> fset_rel A ===> fset_rel B) fimage fimage"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   881
  unfolding fun_rel_def fset_rel_alt_def by simp blast
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   882
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   883
lemma fBall_transfer [transfer_rule]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   884
  "(fset_rel A ===> (A ===> op =) ===> op =) fBall fBall"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   885
  unfolding fset_rel_alt_def fun_rel_def by blast
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   886
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   887
lemma fBex_transfer [transfer_rule]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   888
  "(fset_rel A ===> (A ===> op =) ===> op =) fBex fBex"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   889
  unfolding fset_rel_alt_def fun_rel_def by blast
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   890
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   891
(* FIXME transfer doesn't work here *)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   892
lemma fPow_transfer [transfer_rule]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   893
  "(fset_rel A ===> fset_rel (fset_rel A)) fPow fPow"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   894
  unfolding fun_rel_def
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   895
  using Pow_transfer[unfolded fun_rel_def, rule_format, Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   896
  by blast
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   897
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   898
lemma fset_rel_transfer [transfer_rule]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   899
  "((A ===> B ===> op =) ===> fset_rel A ===> fset_rel B ===> op =)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   900
    fset_rel fset_rel"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   901
  unfolding fun_rel_def
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   902
  using set_rel_transfer[unfolded fun_rel_def,rule_format, Transfer.transferred, where A = A and B = B]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   903
  by simp
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   904
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   905
lemma bind_transfer [transfer_rule]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   906
  "(fset_rel A ===> (A ===> fset_rel B) ===> fset_rel B) fbind fbind"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   907
  using assms unfolding fun_rel_def
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   908
  using bind_transfer[unfolded fun_rel_def, rule_format, Transfer.transferred] by blast
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   909
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   910
text {* Rules requiring bi-unique, bi-total or right-total relations *}
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   911
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   912
lemma fmember_transfer [transfer_rule]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   913
  assumes "bi_unique A"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   914
  shows "(A ===> fset_rel A ===> op =) (op |\<in>|) (op |\<in>|)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   915
  using assms unfolding fun_rel_def fset_rel_alt_def bi_unique_def by metis
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   916
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   917
lemma finter_transfer [transfer_rule]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   918
  assumes "bi_unique A"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   919
  shows "(fset_rel A ===> fset_rel A ===> fset_rel A) finter finter"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   920
  using assms unfolding fun_rel_def
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   921
  using inter_transfer[unfolded fun_rel_def, rule_format, Transfer.transferred] by blast
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   922
53963
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   923
lemma fminus_transfer [transfer_rule]:
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   924
  assumes "bi_unique A"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   925
  shows "(fset_rel A ===> fset_rel A ===> fset_rel A) (op |-|) (op |-|)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   926
  using assms unfolding fun_rel_def
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   927
  using Diff_transfer[unfolded fun_rel_def, rule_format, Transfer.transferred] by blast
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   928
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   929
lemma fsubset_transfer [transfer_rule]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   930
  assumes "bi_unique A"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   931
  shows "(fset_rel A ===> fset_rel A ===> op =) (op |\<subseteq>|) (op |\<subseteq>|)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   932
  using assms unfolding fun_rel_def
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   933
  using subset_transfer[unfolded fun_rel_def, rule_format, Transfer.transferred] by blast
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   934
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   935
lemma fSup_transfer [transfer_rule]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   936
  "bi_unique A \<Longrightarrow> (set_rel (fset_rel A) ===> fset_rel A) Sup Sup"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   937
  using assms unfolding fun_rel_def
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   938
  apply clarify
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   939
  apply transfer'
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   940
  using Sup_fset_transfer[unfolded fun_rel_def] by blast
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   941
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   942
(* FIXME: add right_total_fInf_transfer *)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   943
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   944
lemma fInf_transfer [transfer_rule]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   945
  assumes "bi_unique A" and "bi_total A"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   946
  shows "(set_rel (fset_rel A) ===> fset_rel A) Inf Inf"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   947
  using assms unfolding fun_rel_def
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   948
  apply clarify
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   949
  apply transfer'
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   950
  using Inf_fset_transfer[unfolded fun_rel_def] by blast
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   951
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   952
lemma ffilter_transfer [transfer_rule]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   953
  assumes "bi_unique A"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   954
  shows "((A ===> op=) ===> fset_rel A ===> fset_rel A) ffilter ffilter"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   955
  using assms unfolding fun_rel_def
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   956
  using Lifting_Set.filter_transfer[unfolded fun_rel_def, rule_format, Transfer.transferred] by blast
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   957
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   958
lemma card_transfer [transfer_rule]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   959
  "bi_unique A \<Longrightarrow> (fset_rel A ===> op =) fcard fcard"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   960
  using assms unfolding fun_rel_def
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   961
  using card_transfer[unfolded fun_rel_def, rule_format, Transfer.transferred] by blast
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   962
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   963
end
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   964
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   965
lifting_update fset.lifting
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   966
lifting_forget fset.lifting
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   967
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   968
end