src/HOL/Library/FSet.thy
author haftmann
Sun, 15 Nov 2020 07:17:06 +0000
changeset 72607 feebdaa346e5
parent 72581 de581f98a3a1
child 73832 9db620f007fa
permissions -rw-r--r--
bundles for reflected term syntax
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
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
     4
    Author:     Andrei Popescu, TU Muenchen
53953
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
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60228
diff changeset
     7
section \<open>Type of finite sets defined as a subtype of sets\<close>
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
     8
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
     9
theory FSet
66262
4a2c9d32e7aa finite sets are countable
Lars Hupel <lars.hupel@mytum.de>
parents: 66261
diff changeset
    10
imports Main Countable
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    11
begin
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    12
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60228
diff changeset
    13
subsection \<open>Definition of the type\<close>
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    14
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    15
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
    16
by auto
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    17
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    18
setup_lifting type_definition_fset
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    19
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
    20
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60228
diff changeset
    21
subsection \<open>Basic operations and type class instantiations\<close>
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    22
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    23
(* FIXME transfer and right_total vs. bi_total *)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    24
instantiation fset :: (finite) finite
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    25
begin
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    26
instance by (standard; transfer; simp)
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    27
end
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    28
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    29
instantiation fset :: (type) "{bounded_lattice_bot, distrib_lattice, minus}"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    30
begin
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    31
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
    32
lift_definition bot_fset :: "'a fset" is "{}" parametric empty_transfer by simp
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    33
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
    34
lift_definition less_eq_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" is subset_eq parametric subset_transfer
55565
f663fc1e653b simplify proofs because of the stronger reflexivity prover
kuncar
parents: 55414
diff changeset
    35
  .
53953
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]:
63343
fb5d8a50c641 bundle lifting_syntax;
wenzelm
parents: 63331
diff changeset
    40
  includes lifting_syntax
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
    41
  assumes [transfer_rule]: "bi_unique A"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67091
diff changeset
    42
  shows "((pcr_fset A) ===> (pcr_fset A) ===> (=)) (\<subset>) (<)"
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    43
  unfolding less_fset_def[abs_def] psubset_eq[abs_def] by transfer_prover
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
    44
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    45
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    46
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
    47
  by simp
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    48
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    49
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
    50
  by simp
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    51
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    52
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
    53
  by simp
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    54
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    55
instance
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    56
  by (standard; transfer; auto)+
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    57
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    58
end
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    59
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    60
abbreviation fempty :: "'a fset" ("{||}") where "{||} \<equiv> bot"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    61
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
    62
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
    63
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
    64
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
    65
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
    66
54014
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53969
diff changeset
    67
instantiation fset :: (equal) equal
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53969
diff changeset
    68
begin
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53969
diff changeset
    69
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
    70
instance by intro_classes (auto simp add: equal_fset_def)
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
    71
end
54014
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53969
diff changeset
    72
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    73
instantiation fset :: (type) conditionally_complete_lattice
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    74
begin
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    75
63343
fb5d8a50c641 bundle lifting_syntax;
wenzelm
parents: 63331
diff changeset
    76
context includes lifting_syntax
fb5d8a50c641 bundle lifting_syntax;
wenzelm
parents: 63331
diff changeset
    77
begin
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    78
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    79
lemma right_total_Inf_fset_transfer:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    80
  assumes [transfer_rule]: "bi_unique A" and [transfer_rule]: "right_total A"
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
    81
  shows "(rel_set (rel_set A) ===> rel_set A)
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
    82
    (\<lambda>S. if finite (\<Inter>S \<inter> Collect (Domainp A)) then \<Inter>S \<inter> Collect (Domainp A) else {})
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    83
      (\<lambda>S. if finite (Inf S) then Inf S else {})"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    84
    by transfer_prover
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    85
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    86
lemma Inf_fset_transfer:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    87
  assumes [transfer_rule]: "bi_unique A" and [transfer_rule]: "bi_total A"
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
    88
  shows "(rel_set (rel_set A) ===> rel_set A) (\<lambda>A. if finite (Inf A) then Inf A else {})
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    89
    (\<lambda>A. if finite (Inf A) then Inf A else {})"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    90
  by transfer_prover
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    91
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
    92
lift_definition Inf_fset :: "'a fset set \<Rightarrow> 'a fset" is "\<lambda>A. if finite (Inf A) then Inf A else {}"
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    93
parametric right_total_Inf_fset_transfer Inf_fset_transfer by simp
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    94
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    95
lemma Sup_fset_transfer:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    96
  assumes [transfer_rule]: "bi_unique A"
55938
f20d1db5aa3c renamed 'set_rel' to 'rel_set'
blanchet
parents: 55933
diff changeset
    97
  shows "(rel_set (rel_set A) ===> rel_set A) (\<lambda>A. if finite (Sup A) then Sup A else {})
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
    98
  (\<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
    99
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   100
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
   101
parametric Sup_fset_transfer by simp
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   102
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   103
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
   104
by (auto intro: finite_subset)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   105
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67091
diff changeset
   106
lemma transfer_bdd_below[transfer_rule]: "(rel_set (pcr_fset (=)) ===> (=)) bdd_below bdd_below"
54258
adfc759263ab use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents: 54014
diff changeset
   107
  by auto
adfc759263ab use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents: 54014
diff changeset
   108
63343
fb5d8a50c641 bundle lifting_syntax;
wenzelm
parents: 63331
diff changeset
   109
end
fb5d8a50c641 bundle lifting_syntax;
wenzelm
parents: 63331
diff changeset
   110
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   111
instance
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
   112
proof
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   113
  fix x z :: "'a fset"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   114
  fix X :: "'a fset set"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   115
  {
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
   116
    assume "x \<in> X" "bdd_below X"
56646
360a05d60761 added 'size' of finite sets
blanchet
parents: 56525
diff changeset
   117
    then show "Inf X |\<subseteq>| x" by transfer auto
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   118
  next
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   119
    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
   120
    then show "z |\<subseteq>| Inf X" by transfer (clarsimp, blast)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   121
  next
54258
adfc759263ab use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents: 54014
diff changeset
   122
    assume "x \<in> X" "bdd_above X"
adfc759263ab use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents: 54014
diff changeset
   123
    then obtain z where "x \<in> X" "(\<And>x. x \<in> X \<Longrightarrow> x |\<subseteq>| z)"
adfc759263ab use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents: 54014
diff changeset
   124
      by (auto simp: bdd_above_def)
adfc759263ab use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents: 54014
diff changeset
   125
    then show "x |\<subseteq>| Sup X"
adfc759263ab use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents: 54014
diff changeset
   126
      by transfer (auto intro!: finite_Sup)
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   127
  next
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   128
    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
   129
    then show "Sup X |\<subseteq>| z" by transfer (clarsimp, blast)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   130
  }
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   131
qed
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   132
end
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   133
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
   134
instantiation fset :: (finite) complete_lattice
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   135
begin
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   136
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   137
lift_definition top_fset :: "'a fset" is UNIV parametric right_total_UNIV_transfer UNIV_transfer
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   138
  by simp
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   139
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   140
instance
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   141
  by (standard; transfer; auto)
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   142
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   143
end
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
instantiation fset :: (finite) complete_boolean_algebra
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   146
begin
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   147
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
   148
lift_definition uminus_fset :: "'a fset \<Rightarrow> 'a fset" is uminus
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   149
  parametric right_total_Compl_transfer Compl_transfer by simp
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   150
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   151
instance
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67764
diff changeset
   152
  by (standard; transfer) (simp_all add: Inf_Sup Diff_eq)
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   153
end
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   154
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   155
abbreviation fUNIV :: "'a::finite fset" where "fUNIV \<equiv> top"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   156
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
   157
56646
360a05d60761 added 'size' of finite sets
blanchet
parents: 56525
diff changeset
   158
declare top_fset.rep_eq[simp]
360a05d60761 added 'size' of finite sets
blanchet
parents: 56525
diff changeset
   159
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
   160
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60228
diff changeset
   161
subsection \<open>Other operations\<close>
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   162
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   163
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
   164
  by simp
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
syntax
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   167
  "_insert_fset"     :: "args => 'a fset"  ("{|(_)|}")
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
translations
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   170
  "{|x, xs|}" == "CONST finsert x {|xs|}"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   171
  "{|x|}"     == "CONST finsert x {||}"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   172
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
   173
lift_definition fmember :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<in>|" 50) is Set.member
55565
f663fc1e653b simplify proofs because of the stronger reflexivity prover
kuncar
parents: 55414
diff changeset
   174
  parametric member_transfer .
53953
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
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
   177
63343
fb5d8a50c641 bundle lifting_syntax;
wenzelm
parents: 63331
diff changeset
   178
context includes lifting_syntax
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   179
begin
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
   180
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
   181
lift_definition ffilter :: "('a \<Rightarrow> bool) \<Rightarrow> 'a fset \<Rightarrow> 'a fset" is Set.filter
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   182
  parametric Lifting_Set.filter_transfer unfolding Set.filter_def by simp
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   183
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
   184
lift_definition fPow :: "'a fset \<Rightarrow> 'a fset fset" is Pow parametric Pow_transfer
55732
07906fc6af7a simplify and repair proofs due to df0fda378813
kuncar
parents: 55565
diff changeset
   185
by (simp add: finite_subset)
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   186
55565
f663fc1e653b simplify proofs because of the stronger reflexivity prover
kuncar
parents: 55414
diff changeset
   187
lift_definition fcard :: "'a fset \<Rightarrow> nat" is card parametric card_transfer .
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   188
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
   189
lift_definition fimage :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset" (infixr "|`|" 90) is image
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   190
  parametric image_transfer by simp
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   191
55565
f663fc1e653b simplify proofs because of the stronger reflexivity prover
kuncar
parents: 55414
diff changeset
   192
lift_definition fthe_elem :: "'a fset \<Rightarrow> 'a" is the_elem .
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   193
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
   194
lift_definition fbind :: "'a fset \<Rightarrow> ('a \<Rightarrow> 'b fset) \<Rightarrow> 'b fset" is Set.bind parametric bind_transfer
55738
303123344459 simplify a proof due to 6c95a39348bd
kuncar
parents: 55732
diff changeset
   195
by (simp add: Set.bind_def)
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   196
55732
07906fc6af7a simplify and repair proofs due to df0fda378813
kuncar
parents: 55565
diff changeset
   197
lift_definition ffUnion :: "'a fset fset \<Rightarrow> 'a fset" is Union parametric Union_transfer by simp
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   198
55565
f663fc1e653b simplify proofs because of the stronger reflexivity prover
kuncar
parents: 55414
diff changeset
   199
lift_definition fBall :: "'a fset \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" is Ball parametric Ball_transfer .
f663fc1e653b simplify proofs because of the stronger reflexivity prover
kuncar
parents: 55414
diff changeset
   200
lift_definition fBex :: "'a fset \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" is Bex parametric Bex_transfer .
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   201
55565
f663fc1e653b simplify proofs because of the stronger reflexivity prover
kuncar
parents: 55414
diff changeset
   202
lift_definition ffold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a fset \<Rightarrow> 'b" is Finite_Set.fold .
53963
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   203
63622
7fb02cee1cba some additions to FSet
Lars Hupel <lars.hupel@mytum.de>
parents: 63343
diff changeset
   204
lift_definition fset_of_list :: "'a list \<Rightarrow> 'a fset" is set by (rule finite_set)
7fb02cee1cba some additions to FSet
Lars Hupel <lars.hupel@mytum.de>
parents: 63343
diff changeset
   205
68463
410818a69ee3 material on finite sets and maps
Lars Hupel <lars.hupel@mytum.de>
parents: 67829
diff changeset
   206
lift_definition sorted_list_of_fset :: "'a::linorder fset \<Rightarrow> 'a list" is sorted_list_of_set .
410818a69ee3 material on finite sets and maps
Lars Hupel <lars.hupel@mytum.de>
parents: 67829
diff changeset
   207
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60228
diff changeset
   208
subsection \<open>Transferred lemmas from Set.thy\<close>
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   209
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   210
lemmas fset_eqI = set_eqI[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   211
lemmas fset_eq_iff[no_atp] = set_eq_iff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   212
lemmas fBallI[intro!] = ballI[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   213
lemmas fbspec[dest?] = bspec[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   214
lemmas fBallE[elim] = ballE[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   215
lemmas fBexI[intro] = bexI[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   216
lemmas rev_fBexI[intro?] = rev_bexI[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   217
lemmas fBexCI = bexCI[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   218
lemmas fBexE[elim!] = bexE[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   219
lemmas fBall_triv[simp] = ball_triv[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   220
lemmas fBex_triv[simp] = bex_triv[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   221
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
   222
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
   223
lemmas fBex_one_point1[simp] = bex_one_point1[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   224
lemmas fBex_one_point2[simp] = bex_one_point2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   225
lemmas fBall_one_point1[simp] = ball_one_point1[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   226
lemmas fBall_one_point2[simp] = ball_one_point2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   227
lemmas fBall_conj_distrib = ball_conj_distrib[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   228
lemmas fBex_disj_distrib = bex_disj_distrib[Transfer.transferred]
66264
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   229
lemmas fBall_cong[fundef_cong] = ball_cong[Transfer.transferred]
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   230
lemmas fBex_cong[fundef_cong] = bex_cong[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   231
lemmas fsubsetI[intro!] = subsetI[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   232
lemmas fsubsetD[elim, intro?] = subsetD[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   233
lemmas rev_fsubsetD[no_atp,intro?] = rev_subsetD[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   234
lemmas fsubsetCE[no_atp,elim] = subsetCE[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   235
lemmas fsubset_eq[no_atp] = subset_eq[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   236
lemmas contra_fsubsetD[no_atp] = contra_subsetD[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   237
lemmas fsubset_refl = subset_refl[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   238
lemmas fsubset_trans = subset_trans[Transfer.transferred]
69712
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69700
diff changeset
   239
lemmas fset_rev_mp = rev_subsetD[Transfer.transferred]
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69700
diff changeset
   240
lemmas fset_mp = subsetD[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   241
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
   242
lemmas eq_fmem_trans = eq_mem_trans[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   243
lemmas fsubset_antisym[intro!] = subset_antisym[Transfer.transferred]
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   244
lemmas fequalityD1 = equalityD1[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   245
lemmas fequalityD2 = equalityD2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   246
lemmas fequalityE = equalityE[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   247
lemmas fequalityCE[elim] = equalityCE[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   248
lemmas eqfset_imp_iff = eqset_imp_iff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   249
lemmas eqfelem_imp_iff = eqelem_imp_iff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   250
lemmas fempty_iff[simp] = empty_iff[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   251
lemmas fempty_fsubsetI[iff] = empty_subsetI[Transfer.transferred]
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   252
lemmas equalsffemptyI = equals0I[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   253
lemmas equalsffemptyD = equals0D[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   254
lemmas fBall_fempty[simp] = ball_empty[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   255
lemmas fBex_fempty[simp] = bex_empty[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   256
lemmas fPow_iff[iff] = Pow_iff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   257
lemmas fPowI = PowI[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   258
lemmas fPowD = PowD[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   259
lemmas fPow_bottom = Pow_bottom[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   260
lemmas fPow_top = Pow_top[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   261
lemmas fPow_not_fempty = Pow_not_empty[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   262
lemmas finter_iff[simp] = Int_iff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   263
lemmas finterI[intro!] = IntI[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   264
lemmas finterD1 = IntD1[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   265
lemmas finterD2 = IntD2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   266
lemmas finterE[elim!] = IntE[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   267
lemmas funion_iff[simp] = Un_iff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   268
lemmas funionI1[elim?] = UnI1[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   269
lemmas funionI2[elim?] = UnI2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   270
lemmas funionCI[intro!] = UnCI[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   271
lemmas funionE[elim!] = UnE[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   272
lemmas fminus_iff[simp] = Diff_iff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   273
lemmas fminusI[intro!] = DiffI[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   274
lemmas fminusD1 = DiffD1[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   275
lemmas fminusD2 = DiffD2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   276
lemmas fminusE[elim!] = DiffE[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   277
lemmas finsert_iff[simp] = insert_iff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   278
lemmas finsertI1 = insertI1[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   279
lemmas finsertI2 = insertI2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   280
lemmas finsertE[elim!] = insertE[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   281
lemmas finsertCI[intro!] = insertCI[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   282
lemmas fsubset_finsert_iff = subset_insert_iff[Transfer.transferred]
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   283
lemmas finsert_ident = insert_ident[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   284
lemmas fsingletonI[intro!,no_atp] = singletonI[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   285
lemmas fsingletonD[dest!,no_atp] = singletonD[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   286
lemmas fsingleton_iff = singleton_iff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   287
lemmas fsingleton_inject[dest!] = singleton_inject[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   288
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
   289
lemmas fsingleton_finsert_inj_eq'[iff,no_atp] = singleton_insert_inj_eq'[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   290
lemmas fsubset_fsingletonD = subset_singletonD[Transfer.transferred]
62087
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 62082
diff changeset
   291
lemmas fminus_single_finsert = Diff_single_insert[Transfer.transferred]
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   292
lemmas fdoubleton_eq_iff = doubleton_eq_iff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   293
lemmas funion_fsingleton_iff = Un_singleton_iff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   294
lemmas fsingleton_funion_iff = singleton_Un_iff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   295
lemmas fimage_eqI[simp, intro] = image_eqI[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   296
lemmas fimageI = imageI[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   297
lemmas rev_fimage_eqI = rev_image_eqI[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   298
lemmas fimageE[elim!] = imageE[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   299
lemmas Compr_fimage_eq = Compr_image_eq[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   300
lemmas fimage_funion = image_Un[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   301
lemmas fimage_iff = image_iff[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   302
lemmas fimage_fsubset_iff[no_atp] = image_subset_iff[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   303
lemmas fimage_fsubsetI = image_subsetI[Transfer.transferred]
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   304
lemmas fimage_ident[simp] = image_ident[Transfer.transferred]
62390
842917225d56 more canonical names
nipkow
parents: 62343
diff changeset
   305
lemmas if_split_fmem1 = if_split_mem1[Transfer.transferred]
842917225d56 more canonical names
nipkow
parents: 62343
diff changeset
   306
lemmas if_split_fmem2 = if_split_mem2[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   307
lemmas pfsubsetI[intro!,no_atp] = psubsetI[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   308
lemmas pfsubsetE[elim!,no_atp] = psubsetE[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   309
lemmas pfsubset_finsert_iff = psubset_insert_iff[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   310
lemmas pfsubset_eq = psubset_eq[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   311
lemmas pfsubset_imp_fsubset = psubset_imp_subset[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   312
lemmas pfsubset_trans = psubset_trans[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   313
lemmas pfsubsetD = psubsetD[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   314
lemmas pfsubset_fsubset_trans = psubset_subset_trans[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   315
lemmas fsubset_pfsubset_trans = subset_psubset_trans[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   316
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
   317
lemmas fimage_fPow_mono = image_Pow_mono[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   318
lemmas fimage_fPow_surj = image_Pow_surj[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   319
lemmas fsubset_finsertI = subset_insertI[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   320
lemmas fsubset_finsertI2 = subset_insertI2[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   321
lemmas fsubset_finsert = subset_insert[Transfer.transferred]
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   322
lemmas funion_upper1 = Un_upper1[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   323
lemmas funion_upper2 = Un_upper2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   324
lemmas funion_least = Un_least[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   325
lemmas finter_lower1 = Int_lower1[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   326
lemmas finter_lower2 = Int_lower2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   327
lemmas finter_greatest = Int_greatest[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   328
lemmas fminus_fsubset = Diff_subset[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   329
lemmas fminus_fsubset_conv = Diff_subset_conv[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   330
lemmas fsubset_fempty[simp] = subset_empty[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   331
lemmas not_pfsubset_fempty[iff] = not_psubset_empty[Transfer.transferred]
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   332
lemmas finsert_is_funion = insert_is_Un[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   333
lemmas finsert_not_fempty[simp] = insert_not_empty[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   334
lemmas fempty_not_finsert = empty_not_insert[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   335
lemmas finsert_absorb = insert_absorb[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   336
lemmas finsert_absorb2[simp] = insert_absorb2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   337
lemmas finsert_commute = insert_commute[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   338
lemmas finsert_fsubset[simp] = insert_subset[Transfer.transferred]
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   339
lemmas finsert_inter_finsert[simp] = insert_inter_insert[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   340
lemmas finsert_disjoint[simp,no_atp] = insert_disjoint[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   341
lemmas disjoint_finsert[simp,no_atp] = disjoint_insert[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   342
lemmas fimage_fempty[simp] = image_empty[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   343
lemmas fimage_finsert[simp] = image_insert[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   344
lemmas fimage_constant = image_constant[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   345
lemmas fimage_constant_conv = image_constant_conv[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   346
lemmas fimage_fimage = image_image[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   347
lemmas finsert_fimage[simp] = insert_image[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   348
lemmas fimage_is_fempty[iff] = image_is_empty[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   349
lemmas fempty_is_fimage[iff] = empty_is_image[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   350
lemmas fimage_cong = image_cong[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   351
lemmas fimage_finter_fsubset = image_Int_subset[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   352
lemmas fimage_fminus_fsubset = image_diff_subset[Transfer.transferred]
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   353
lemmas finter_absorb = Int_absorb[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   354
lemmas finter_left_absorb = Int_left_absorb[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   355
lemmas finter_commute = Int_commute[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   356
lemmas finter_left_commute = Int_left_commute[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   357
lemmas finter_assoc = Int_assoc[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   358
lemmas finter_ac = Int_ac[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   359
lemmas finter_absorb1 = Int_absorb1[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   360
lemmas finter_absorb2 = Int_absorb2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   361
lemmas finter_fempty_left = Int_empty_left[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   362
lemmas finter_fempty_right = Int_empty_right[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   363
lemmas disjoint_iff_fnot_equal = disjoint_iff_not_equal[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   364
lemmas finter_funion_distrib = Int_Un_distrib[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   365
lemmas finter_funion_distrib2 = Int_Un_distrib2[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   366
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
   367
lemmas funion_absorb = Un_absorb[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   368
lemmas funion_left_absorb = Un_left_absorb[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   369
lemmas funion_commute = Un_commute[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   370
lemmas funion_left_commute = Un_left_commute[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   371
lemmas funion_assoc = Un_assoc[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   372
lemmas funion_ac = Un_ac[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   373
lemmas funion_absorb1 = Un_absorb1[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   374
lemmas funion_absorb2 = Un_absorb2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   375
lemmas funion_fempty_left = Un_empty_left[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   376
lemmas funion_fempty_right = Un_empty_right[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   377
lemmas funion_finsert_left[simp] = Un_insert_left[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   378
lemmas funion_finsert_right[simp] = Un_insert_right[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   379
lemmas finter_finsert_left = Int_insert_left[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   380
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
   381
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
   382
lemmas finter_finsert_right = Int_insert_right[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   383
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
   384
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
   385
lemmas funion_finter_distrib = Un_Int_distrib[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   386
lemmas funion_finter_distrib2 = Un_Int_distrib2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   387
lemmas funion_finter_crazy = Un_Int_crazy[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   388
lemmas fsubset_funion_eq = subset_Un_eq[Transfer.transferred]
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   389
lemmas funion_fempty[iff] = Un_empty[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   390
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
   391
lemmas funion_fminus_finter = Un_Diff_Int[Transfer.transferred]
68463
410818a69ee3 material on finite sets and maps
Lars Hupel <lars.hupel@mytum.de>
parents: 67829
diff changeset
   392
lemmas ffunion_empty[simp] = Union_empty[Transfer.transferred]
410818a69ee3 material on finite sets and maps
Lars Hupel <lars.hupel@mytum.de>
parents: 67829
diff changeset
   393
lemmas ffunion_mono = Union_mono[Transfer.transferred]
410818a69ee3 material on finite sets and maps
Lars Hupel <lars.hupel@mytum.de>
parents: 67829
diff changeset
   394
lemmas ffunion_insert[simp] = Union_insert[Transfer.transferred]
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   395
lemmas fminus_finter2 = Diff_Int2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   396
lemmas funion_finter_assoc_eq = Un_Int_assoc_eq[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   397
lemmas fBall_funion = ball_Un[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   398
lemmas fBex_funion = bex_Un[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   399
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
   400
lemmas fminus_cancel[simp] = Diff_cancel[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   401
lemmas fminus_idemp[simp] = Diff_idemp[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   402
lemmas fminus_triv = Diff_triv[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   403
lemmas fempty_fminus[simp] = empty_Diff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   404
lemmas fminus_fempty[simp] = Diff_empty[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   405
lemmas fminus_finsertffempty[simp,no_atp] = Diff_insert0[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   406
lemmas fminus_finsert = Diff_insert[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   407
lemmas fminus_finsert2 = Diff_insert2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   408
lemmas finsert_fminus_if = insert_Diff_if[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   409
lemmas finsert_fminus1[simp] = insert_Diff1[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   410
lemmas finsert_fminus_single[simp] = insert_Diff_single[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   411
lemmas finsert_fminus = insert_Diff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   412
lemmas fminus_finsert_absorb = Diff_insert_absorb[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   413
lemmas fminus_disjoint[simp] = Diff_disjoint[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   414
lemmas fminus_partition = Diff_partition[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   415
lemmas double_fminus = double_diff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   416
lemmas funion_fminus_cancel[simp] = Un_Diff_cancel[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   417
lemmas funion_fminus_cancel2[simp] = Un_Diff_cancel2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   418
lemmas fminus_funion = Diff_Un[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   419
lemmas fminus_finter = Diff_Int[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   420
lemmas funion_fminus = Un_Diff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   421
lemmas finter_fminus = Int_Diff[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   422
lemmas fminus_finter_distrib = Diff_Int_distrib[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   423
lemmas fminus_finter_distrib2 = Diff_Int_distrib2[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   424
lemmas fUNIV_bool[no_atp] = UNIV_bool[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   425
lemmas fPow_fempty[simp] = Pow_empty[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   426
lemmas fPow_finsert = Pow_insert[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   427
lemmas funion_fPow_fsubset = Un_Pow_subset[Transfer.transferred]
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   428
lemmas fPow_finter_eq[simp] = Pow_Int_eq[Transfer.transferred]
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   429
lemmas fset_eq_fsubset = set_eq_subset[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   430
lemmas fsubset_iff[no_atp] = subset_iff[Transfer.transferred]
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   431
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
   432
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
   433
lemmas ex_fin_conv = ex_in_conv[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   434
lemmas fimage_mono = image_mono[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   435
lemmas fPow_mono = Pow_mono[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   436
lemmas finsert_mono = insert_mono[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   437
lemmas funion_mono = Un_mono[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   438
lemmas finter_mono = Int_mono[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   439
lemmas fminus_mono = Diff_mono[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   440
lemmas fin_mono = in_mono[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   441
lemmas fthe_felem_eq[simp] = the_elem_eq[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   442
lemmas fLeast_mono = Least_mono[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   443
lemmas fbind_fbind = bind_bind[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   444
lemmas fempty_fbind[simp] = empty_bind[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   445
lemmas nonfempty_fbind_const = nonempty_bind_const[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   446
lemmas fbind_const = bind_const[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   447
lemmas ffmember_filter[simp] = member_filter[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   448
lemmas fequalityI = equalityI[Transfer.transferred]
63622
7fb02cee1cba some additions to FSet
Lars Hupel <lars.hupel@mytum.de>
parents: 63343
diff changeset
   449
lemmas fset_of_list_simps[simp] = set_simps[Transfer.transferred]
7fb02cee1cba some additions to FSet
Lars Hupel <lars.hupel@mytum.de>
parents: 63343
diff changeset
   450
lemmas fset_of_list_append[simp] = set_append[Transfer.transferred]
7fb02cee1cba some additions to FSet
Lars Hupel <lars.hupel@mytum.de>
parents: 63343
diff changeset
   451
lemmas fset_of_list_rev[simp] = set_rev[Transfer.transferred]
7fb02cee1cba some additions to FSet
Lars Hupel <lars.hupel@mytum.de>
parents: 63343
diff changeset
   452
lemmas fset_of_list_map[simp] = set_map[Transfer.transferred]
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   453
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
   454
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60228
diff changeset
   455
subsection \<open>Additional lemmas\<close>
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   456
66264
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   457
subsubsection \<open>\<open>ffUnion\<close>\<close>
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   458
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   459
lemmas ffUnion_funion_distrib[simp] = Union_Un_distrib[Transfer.transferred]
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   460
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   461
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   462
subsubsection \<open>\<open>fbind\<close>\<close>
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   463
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   464
lemma fbind_cong[fundef_cong]: "A = B \<Longrightarrow> (\<And>x. x |\<in>| B \<Longrightarrow> f x = g x) \<Longrightarrow> fbind A f = fbind B g"
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   465
by transfer force
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   466
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   467
61585
a9599d3d7610 isabelle update_cartouches -c -t;
wenzelm
parents: 61424
diff changeset
   468
subsubsection \<open>\<open>fsingleton\<close>\<close>
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   469
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   470
lemmas fsingletonE = fsingletonD [elim_format]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   471
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
   472
61585
a9599d3d7610 isabelle update_cartouches -c -t;
wenzelm
parents: 61424
diff changeset
   473
subsubsection \<open>\<open>femepty\<close>\<close>
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   474
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   475
lemma fempty_ffilter[simp]: "ffilter (\<lambda>_. False) A = {||}"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   476
by transfer auto
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
(* FIXME, transferred doesn't work here *)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   479
lemma femptyE [elim!]: "a |\<in>| {||} \<Longrightarrow> P"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   480
  by simp
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   481
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
   482
61585
a9599d3d7610 isabelle update_cartouches -c -t;
wenzelm
parents: 61424
diff changeset
   483
subsubsection \<open>\<open>fset\<close>\<close>
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   484
53963
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   485
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
   486
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
   487
lemma finite_fset [simp]:
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   488
  shows "finite (fset S)"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   489
  by transfer simp
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   490
53963
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   491
lemmas fset_cong = fset_inject
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
lemma filter_fset [simp]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   494
  shows "fset (ffilter P xs) = Collect P \<inter> fset xs"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   495
  by transfer auto
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   496
53963
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   497
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
   498
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   499
lemmas inter_fset[simp] = inf_fset.rep_eq
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   500
53963
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   501
lemmas union_fset[simp] = sup_fset.rep_eq
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   502
53963
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   503
lemmas minus_fset[simp] = minus_fset.rep_eq
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   504
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
   505
63622
7fb02cee1cba some additions to FSet
Lars Hupel <lars.hupel@mytum.de>
parents: 63343
diff changeset
   506
subsubsection \<open>\<open>ffilter\<close>\<close>
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   507
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
   508
lemma subset_ffilter:
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   509
  "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
   510
  by transfer auto
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   511
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
   512
lemma eq_ffilter:
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   513
  "(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
   514
  by transfer auto
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   515
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   516
lemma pfsubset_ffilter:
67091
1393c2340eec more symbols;
wenzelm
parents: 66292
diff changeset
   517
  "(\<And>x. x |\<in>| A \<Longrightarrow> P x \<Longrightarrow> Q x) \<Longrightarrow> (x |\<in>| A \<and> \<not> P x \<and> Q x) \<Longrightarrow>
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   518
    ffilter P A |\<subset>| ffilter Q A"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   519
  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
   520
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
   521
63622
7fb02cee1cba some additions to FSet
Lars Hupel <lars.hupel@mytum.de>
parents: 63343
diff changeset
   522
subsubsection \<open>\<open>fset_of_list\<close>\<close>
7fb02cee1cba some additions to FSet
Lars Hupel <lars.hupel@mytum.de>
parents: 63343
diff changeset
   523
7fb02cee1cba some additions to FSet
Lars Hupel <lars.hupel@mytum.de>
parents: 63343
diff changeset
   524
lemma fset_of_list_filter[simp]:
7fb02cee1cba some additions to FSet
Lars Hupel <lars.hupel@mytum.de>
parents: 63343
diff changeset
   525
  "fset_of_list (filter P xs) = ffilter P (fset_of_list xs)"
7fb02cee1cba some additions to FSet
Lars Hupel <lars.hupel@mytum.de>
parents: 63343
diff changeset
   526
  by transfer (auto simp: Set.filter_def)
7fb02cee1cba some additions to FSet
Lars Hupel <lars.hupel@mytum.de>
parents: 63343
diff changeset
   527
7fb02cee1cba some additions to FSet
Lars Hupel <lars.hupel@mytum.de>
parents: 63343
diff changeset
   528
lemma fset_of_list_subset[intro]:
7fb02cee1cba some additions to FSet
Lars Hupel <lars.hupel@mytum.de>
parents: 63343
diff changeset
   529
  "set xs \<subseteq> set ys \<Longrightarrow> fset_of_list xs |\<subseteq>| fset_of_list ys"
7fb02cee1cba some additions to FSet
Lars Hupel <lars.hupel@mytum.de>
parents: 63343
diff changeset
   530
  by transfer simp
7fb02cee1cba some additions to FSet
Lars Hupel <lars.hupel@mytum.de>
parents: 63343
diff changeset
   531
7fb02cee1cba some additions to FSet
Lars Hupel <lars.hupel@mytum.de>
parents: 63343
diff changeset
   532
lemma fset_of_list_elem: "(x |\<in>| fset_of_list xs) \<longleftrightarrow> (x \<in> set xs)"
7fb02cee1cba some additions to FSet
Lars Hupel <lars.hupel@mytum.de>
parents: 63343
diff changeset
   533
  by transfer simp
7fb02cee1cba some additions to FSet
Lars Hupel <lars.hupel@mytum.de>
parents: 63343
diff changeset
   534
7fb02cee1cba some additions to FSet
Lars Hupel <lars.hupel@mytum.de>
parents: 63343
diff changeset
   535
61585
a9599d3d7610 isabelle update_cartouches -c -t;
wenzelm
parents: 61424
diff changeset
   536
subsubsection \<open>\<open>finsert\<close>\<close>
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   537
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   538
(* FIXME, transferred doesn't work here *)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   539
lemma set_finsert:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   540
  assumes "x |\<in>| A"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   541
  obtains B where "A = finsert x B" and "x |\<notin>| B"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   542
using assms by transfer (metis Set.set_insert finite_insert)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   543
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   544
lemma mk_disjoint_finsert: "a |\<in>| A \<Longrightarrow> \<exists>B. A = finsert a B \<and> a |\<notin>| B"
63649
e690d6f2185b tuned proofs;
wenzelm
parents: 63622
diff changeset
   545
  by (rule exI [where x = "A |-| {|a|}"]) blast
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   546
66264
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   547
lemma finsert_eq_iff:
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   548
  assumes "a |\<notin>| A" and "b |\<notin>| B"
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   549
  shows "(finsert a A = finsert b B) =
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   550
    (if a = b then A = B else \<exists>C. A = finsert b C \<and> b |\<notin>| C \<and> B = finsert a C \<and> a |\<notin>| C)"
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   551
  using assms by transfer (force simp: insert_eq_iff)
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   552
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
   553
61585
a9599d3d7610 isabelle update_cartouches -c -t;
wenzelm
parents: 61424
diff changeset
   554
subsubsection \<open>\<open>fimage\<close>\<close>
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   555
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   556
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
   557
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
   558
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
   559
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60228
diff changeset
   560
subsubsection \<open>bounded quantification\<close>
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   561
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   562
lemma bex_simps [simp, no_atp]:
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
   563
  "\<And>A P Q. fBex A (\<lambda>x. P x \<and> Q) = (fBex A P \<and> Q)"
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   564
  "\<And>A P Q. fBex A (\<lambda>x. P \<and> Q x) = (P \<and> fBex A Q)"
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
   565
  "\<And>P. fBex {||} P = False"
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   566
  "\<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
   567
  "\<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
   568
  "\<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
   569
by auto
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   570
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   571
lemma ball_simps [simp, no_atp]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   572
  "\<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
   573
  "\<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
   574
  "\<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
   575
  "\<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
   576
  "\<And>P. fBall {||} P = True"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   577
  "\<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
   578
  "\<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
   579
  "\<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
   580
by auto
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   581
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   582
lemma atomize_fBall:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   583
    "(\<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
   584
apply (simp only: atomize_all atomize_imp)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   585
apply (rule equal_intr_rule)
63622
7fb02cee1cba some additions to FSet
Lars Hupel <lars.hupel@mytum.de>
parents: 63343
diff changeset
   586
  by (transfer, simp)+
7fb02cee1cba some additions to FSet
Lars Hupel <lars.hupel@mytum.de>
parents: 63343
diff changeset
   587
7fb02cee1cba some additions to FSet
Lars Hupel <lars.hupel@mytum.de>
parents: 63343
diff changeset
   588
lemma fBall_mono[mono]: "P \<le> Q \<Longrightarrow> fBall S P \<le> fBall S Q"
7fb02cee1cba some additions to FSet
Lars Hupel <lars.hupel@mytum.de>
parents: 63343
diff changeset
   589
by auto
7fb02cee1cba some additions to FSet
Lars Hupel <lars.hupel@mytum.de>
parents: 63343
diff changeset
   590
68463
410818a69ee3 material on finite sets and maps
Lars Hupel <lars.hupel@mytum.de>
parents: 67829
diff changeset
   591
lemma fBex_mono[mono]: "P \<le> Q \<Longrightarrow> fBex S P \<le> fBex S Q"
410818a69ee3 material on finite sets and maps
Lars Hupel <lars.hupel@mytum.de>
parents: 67829
diff changeset
   592
by auto
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   593
53963
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   594
end
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   595
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
   596
61585
a9599d3d7610 isabelle update_cartouches -c -t;
wenzelm
parents: 61424
diff changeset
   597
subsubsection \<open>\<open>fcard\<close>\<close>
53963
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   598
53964
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   599
(* FIXME: improve transferred to handle bounded meta quantification *)
ac0e4ca891f9 tuned names
kuncar
parents: 53963
diff changeset
   600
53963
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   601
lemma fcard_fempty:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   602
  "fcard {||} = 0"
72302
d7d90ed4c74e fixed some remarkably ugly proofs
paulson <lp15@cam.ac.uk>
parents: 69712
diff changeset
   603
  by transfer (rule card.empty)
53963
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   604
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   605
lemma fcard_finsert_disjoint:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   606
  "x |\<notin>| A \<Longrightarrow> fcard (finsert x A) = Suc (fcard A)"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   607
  by transfer (rule card_insert_disjoint)
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_finsert_if:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   610
  "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
   611
  by transfer (rule card_insert_if)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   612
66265
a51e72d79670 card_0_eq ~> fcard_0_eq
Lars Hupel <lars.hupel@mytum.de>
parents: 66264
diff changeset
   613
lemma fcard_0_eq [simp, no_atp]:
53963
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   614
  "fcard A = 0 \<longleftrightarrow> A = {||}"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   615
  by transfer (rule card_0_eq)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   616
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   617
lemma fcard_Suc_fminus1:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   618
  "x |\<in>| A \<Longrightarrow> Suc (fcard (A |-| {|x|})) = fcard A"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   619
  by transfer (rule card_Suc_Diff1)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   620
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   621
lemma fcard_fminus_fsingleton:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   622
  "x |\<in>| A \<Longrightarrow> fcard (A |-| {|x|}) = fcard A - 1"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   623
  by transfer (rule card_Diff_singleton)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   624
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   625
lemma fcard_fminus_fsingleton_if:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   626
  "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
   627
  by transfer (rule card_Diff_singleton_if)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   628
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   629
lemma fcard_fminus_finsert[simp]:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   630
  assumes "a |\<in>| A" and "a |\<notin>| B"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   631
  shows "fcard (A |-| finsert a B) = fcard (A |-| B) - 1"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   632
using assms by transfer (rule card_Diff_insert)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   633
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   634
lemma fcard_finsert: "fcard (finsert x A) = Suc (fcard (A |-| {|x|}))"
72302
d7d90ed4c74e fixed some remarkably ugly proofs
paulson <lp15@cam.ac.uk>
parents: 69712
diff changeset
   635
by transfer (rule card.insert_remove)
53963
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   636
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   637
lemma fcard_finsert_le: "fcard A \<le> fcard (finsert x A)"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   638
by transfer (rule card_insert_le)
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 fcard_mono:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   641
  "A |\<subseteq>| B \<Longrightarrow> fcard A \<le> fcard B"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   642
by transfer (rule card_mono)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   643
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   644
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
   645
by transfer (rule card_seteq)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   646
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   647
lemma pfsubset_fcard_mono: "A |\<subset>| B \<Longrightarrow> fcard A < fcard B"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   648
by transfer (rule psubset_card_mono)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   649
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
   650
lemma fcard_funion_finter:
53963
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   651
  "fcard A + fcard B = fcard (A |\<union>| B) + fcard (A |\<inter>| B)"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   652
by transfer (rule card_Un_Int)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   653
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   654
lemma fcard_funion_disjoint:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   655
  "A |\<inter>| B = {||} \<Longrightarrow> fcard (A |\<union>| B) = fcard A + fcard B"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   656
by transfer (rule card_Un_disjoint)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   657
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   658
lemma fcard_funion_fsubset:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   659
  "B |\<subseteq>| A \<Longrightarrow> fcard (A |-| B) = fcard A - fcard B"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   660
by transfer (rule card_Diff_subset)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   661
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   662
lemma diff_fcard_le_fcard_fminus:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   663
  "fcard A - fcard B \<le> fcard(A |-| B)"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   664
by transfer (rule diff_card_le_card_Diff)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   665
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   666
lemma fcard_fminus1_less: "x |\<in>| A \<Longrightarrow> fcard (A |-| {|x|}) < fcard A"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   667
by transfer (rule card_Diff1_less)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   668
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   669
lemma fcard_fminus2_less:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   670
  "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
   671
by transfer (rule card_Diff2_less)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   672
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   673
lemma fcard_fminus1_le: "fcard (A |-| {|x|}) \<le> fcard A"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   674
by transfer (rule card_Diff1_le)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   675
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   676
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
   677
by transfer (rule card_psubset)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   678
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
   679
68463
410818a69ee3 material on finite sets and maps
Lars Hupel <lars.hupel@mytum.de>
parents: 67829
diff changeset
   680
subsubsection \<open>\<open>sorted_list_of_fset\<close>\<close>
410818a69ee3 material on finite sets and maps
Lars Hupel <lars.hupel@mytum.de>
parents: 67829
diff changeset
   681
410818a69ee3 material on finite sets and maps
Lars Hupel <lars.hupel@mytum.de>
parents: 67829
diff changeset
   682
lemma sorted_list_of_fset_simps[simp]:
410818a69ee3 material on finite sets and maps
Lars Hupel <lars.hupel@mytum.de>
parents: 67829
diff changeset
   683
  "set (sorted_list_of_fset S) = fset S"
410818a69ee3 material on finite sets and maps
Lars Hupel <lars.hupel@mytum.de>
parents: 67829
diff changeset
   684
  "fset_of_list (sorted_list_of_fset S) = S"
410818a69ee3 material on finite sets and maps
Lars Hupel <lars.hupel@mytum.de>
parents: 67829
diff changeset
   685
by (transfer, simp)+
410818a69ee3 material on finite sets and maps
Lars Hupel <lars.hupel@mytum.de>
parents: 67829
diff changeset
   686
410818a69ee3 material on finite sets and maps
Lars Hupel <lars.hupel@mytum.de>
parents: 67829
diff changeset
   687
61585
a9599d3d7610 isabelle update_cartouches -c -t;
wenzelm
parents: 61424
diff changeset
   688
subsubsection \<open>\<open>ffold\<close>\<close>
53963
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   689
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   690
(* FIXME: improve transferred to handle bounded meta quantification *)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   691
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   692
context comp_fun_commute
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   693
begin
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   694
  lemmas ffold_empty[simp] = fold_empty[Transfer.transferred]
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   695
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   696
  lemma ffold_finsert [simp]:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   697
    assumes "x |\<notin>| A"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   698
    shows "ffold f z (finsert x A) = f x (ffold f z A)"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   699
    using assms by (transfer fixing: f) (rule fold_insert)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   700
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   701
  lemma ffold_fun_left_comm:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   702
    "f x (ffold f z A) = ffold f (f x z) A"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   703
    by (transfer fixing: f) (rule fold_fun_left_comm)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   704
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   705
  lemma ffold_finsert2:
56646
360a05d60761 added 'size' of finite sets
blanchet
parents: 56525
diff changeset
   706
    "x |\<notin>| A \<Longrightarrow> ffold f z (finsert x A) = ffold f (f x z) A"
53963
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   707
    by (transfer fixing: f) (rule fold_insert2)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   708
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   709
  lemma ffold_rec:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   710
    assumes "x |\<in>| A"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   711
    shows "ffold f z A = f x (ffold f z (A |-| {|x|}))"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   712
    using assms by (transfer fixing: f) (rule fold_rec)
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
   713
53963
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   714
  lemma ffold_finsert_fremove:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   715
    "ffold f z (finsert x A) = f x (ffold f z (A |-| {|x|}))"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   716
     by (transfer fixing: f) (rule fold_insert_remove)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   717
end
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   718
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   719
lemma ffold_fimage:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   720
  assumes "inj_on g (fset A)"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   721
  shows "ffold f z (g |`| A) = ffold (f \<circ> g) z A"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   722
using assms by transfer' (rule fold_image)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   723
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   724
lemma ffold_cong:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   725
  assumes "comp_fun_commute f" "comp_fun_commute g"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   726
  "\<And>x. x |\<in>| A \<Longrightarrow> f x = g x"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   727
    and "s = t" and "A = B"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   728
  shows "ffold f s A = ffold g t B"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   729
using assms by transfer (metis Finite_Set.fold_cong)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   730
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   731
context comp_fun_idem
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   732
begin
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   733
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   734
  lemma ffold_finsert_idem:
56646
360a05d60761 added 'size' of finite sets
blanchet
parents: 56525
diff changeset
   735
    "ffold f z (finsert x A) = f x (ffold f z A)"
53963
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   736
    by (transfer fixing: f) (rule fold_insert_idem)
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
   737
53963
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   738
  declare ffold_finsert [simp del] ffold_finsert_idem [simp]
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
   739
53963
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   740
  lemma ffold_finsert_idem2:
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   741
    "ffold f z (finsert x A) = ffold f (f x z) A"
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   742
    by (transfer fixing: f) (rule fold_insert_idem2)
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   743
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   744
end
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
   745
66292
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   746
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   747
subsubsection \<open>Group operations\<close>
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   748
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   749
locale comm_monoid_fset = comm_monoid
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   750
begin
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   751
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   752
sublocale set: comm_monoid_set ..
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   753
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   754
lift_definition F :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b fset \<Rightarrow> 'a" is set.F .
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   755
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   756
lemmas cong[fundef_cong] = set.cong[Transfer.transferred]
66261
fb6efe575c6d lift sum to finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 64267
diff changeset
   757
69654
bc758f4f09e5 uniform naming
nipkow
parents: 69593
diff changeset
   758
lemma cong_simp[cong]:
69164
74f1b0f10b2b uniform naming of strong congruence rules
nipkow
parents: 68463
diff changeset
   759
  "\<lbrakk> A = B;  \<And>x. x |\<in>| B =simp=> g x = h x \<rbrakk> \<Longrightarrow> F g A = F h B"
74f1b0f10b2b uniform naming of strong congruence rules
nipkow
parents: 68463
diff changeset
   760
unfolding simp_implies_def by (auto cong: cong)
66292
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   761
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   762
end
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   763
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   764
context comm_monoid_add begin
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   765
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   766
sublocale fsum: comm_monoid_fset plus 0
67764
0f8cb5568b63 Drop rewrites after defines in interpretations.
ballarin
parents: 67408
diff changeset
   767
  rewrites "comm_monoid_set.F plus 0 = sum"
66292
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   768
  defines fsum = fsum.F
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   769
proof -
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67091
diff changeset
   770
  show "comm_monoid_fset (+) 0" by standard
66292
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   771
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67091
diff changeset
   772
  show "comm_monoid_set.F (+) 0 = sum" unfolding sum_def ..
66292
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   773
qed
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   774
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   775
end
66261
fb6efe575c6d lift sum to finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 64267
diff changeset
   776
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
   777
66264
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   778
subsubsection \<open>Semilattice operations\<close>
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   779
66292
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   780
locale semilattice_fset = semilattice
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   781
begin
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   782
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   783
sublocale set: semilattice_set ..
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   784
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   785
lift_definition F :: "'a fset \<Rightarrow> 'a" is set.F .
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   786
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   787
lemma eq_fold: "F (finsert x A) = ffold f x A"
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   788
  by transfer (rule set.eq_fold)
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   789
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   790
lemma singleton [simp]: "F {|x|} = x"
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   791
  by transfer (rule set.singleton)
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   792
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   793
lemma insert_not_elem: "x |\<notin>| A \<Longrightarrow> A \<noteq> {||} \<Longrightarrow> F (finsert x A) = x \<^bold>* F A"
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   794
  by transfer (rule set.insert_not_elem)
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   795
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   796
lemma in_idem: "x |\<in>| A \<Longrightarrow> x \<^bold>* F A = F A"
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   797
  by transfer (rule set.in_idem)
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   798
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   799
lemma insert [simp]: "A \<noteq> {||} \<Longrightarrow> F (finsert x A) = x \<^bold>* F A"
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   800
  by transfer (rule set.insert)
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   801
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   802
end
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   803
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   804
locale semilattice_order_fset = binary?: semilattice_order + semilattice_fset
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   805
begin
66264
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   806
66292
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   807
end
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   808
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   809
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   810
context linorder begin
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   811
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   812
sublocale fMin: semilattice_order_fset min less_eq less
67764
0f8cb5568b63 Drop rewrites after defines in interpretations.
ballarin
parents: 67408
diff changeset
   813
  rewrites "semilattice_set.F min = Min"
66292
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   814
  defines fMin = fMin.F
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   815
proof -
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67091
diff changeset
   816
  show "semilattice_order_fset min (\<le>) (<)" by standard
66292
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   817
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   818
  show "semilattice_set.F min = Min" unfolding Min_def ..
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   819
qed
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   820
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   821
sublocale fMax: semilattice_order_fset max greater_eq greater
67764
0f8cb5568b63 Drop rewrites after defines in interpretations.
ballarin
parents: 67408
diff changeset
   822
  rewrites "semilattice_set.F max = Max"
66292
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   823
  defines fMax = fMax.F
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   824
proof -
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67091
diff changeset
   825
  show "semilattice_order_fset max (\<ge>) (>)"
66292
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   826
    by standard
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   827
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   828
  show "semilattice_set.F max = Max"
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   829
    unfolding Max_def ..
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   830
qed
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   831
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   832
end
66264
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   833
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   834
lemma mono_fMax_commute: "mono f \<Longrightarrow> A \<noteq> {||} \<Longrightarrow> f (fMax A) = fMax (f |`| A)"
66292
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   835
  by transfer (rule mono_Max_commute)
66264
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   836
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   837
lemma mono_fMin_commute: "mono f \<Longrightarrow> A \<noteq> {||} \<Longrightarrow> f (fMin A) = fMin (f |`| A)"
66292
9930f4cf6c7a improve setup for fMin/fMax/fsum; courtesy of Ondřej Kunčar & Florian Haftmann
Lars Hupel <lars.hupel@mytum.de>
parents: 66265
diff changeset
   838
  by transfer (rule mono_Min_commute)
66264
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   839
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   840
lemma fMax_in[simp]: "A \<noteq> {||} \<Longrightarrow> fMax A |\<in>| A"
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   841
  by transfer (rule Max_in)
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   842
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   843
lemma fMin_in[simp]: "A \<noteq> {||} \<Longrightarrow> fMin A |\<in>| A"
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   844
  by transfer (rule Min_in)
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   845
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   846
lemma fMax_ge[simp]: "x |\<in>| A \<Longrightarrow> x \<le> fMax A"
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   847
  by transfer (rule Max_ge)
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   848
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   849
lemma fMin_le[simp]: "x |\<in>| A \<Longrightarrow> fMin A \<le> x"
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   850
  by transfer (rule Min_le)
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   851
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   852
lemma fMax_eqI: "(\<And>y. y |\<in>| A \<Longrightarrow> y \<le> x) \<Longrightarrow> x |\<in>| A \<Longrightarrow> fMax A = x"
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   853
  by transfer (rule Max_eqI)
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   854
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   855
lemma fMin_eqI: "(\<And>y. y |\<in>| A \<Longrightarrow> x \<le> y) \<Longrightarrow> x |\<in>| A \<Longrightarrow> fMin A = x"
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   856
  by transfer (rule Min_eqI)
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   857
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   858
lemma fMax_finsert[simp]: "fMax (finsert x A) = (if A = {||} then x else max x (fMax A))"
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   859
  by transfer simp
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   860
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   861
lemma fMin_finsert[simp]: "fMin (finsert x A) = (if A = {||} then x else min x (fMin A))"
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   862
  by transfer simp
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   863
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   864
context linorder begin
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   865
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   866
lemma fset_linorder_max_induct[case_names fempty finsert]:
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   867
  assumes "P {||}"
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   868
  and     "\<And>x S. \<lbrakk>\<forall>y. y |\<in>| S \<longrightarrow> y < x; P S\<rbrakk> \<Longrightarrow> P (finsert x S)"
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   869
  shows "P S"
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   870
proof -
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   871
  (* FIXME transfer and right_total vs. bi_total *)
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   872
  note Domainp_forall_transfer[transfer_rule]
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   873
  show ?thesis
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   874
  using assms by (transfer fixing: less) (auto intro: finite_linorder_max_induct)
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   875
qed
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   876
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   877
lemma fset_linorder_min_induct[case_names fempty finsert]:
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   878
  assumes "P {||}"
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   879
  and     "\<And>x S. \<lbrakk>\<forall>y. y |\<in>| S \<longrightarrow> y > x; P S\<rbrakk> \<Longrightarrow> P (finsert x S)"
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   880
  shows "P S"
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   881
proof -
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   882
  (* FIXME transfer and right_total vs. bi_total *)
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   883
  note Domainp_forall_transfer[transfer_rule]
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   884
  show ?thesis
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   885
  using assms by (transfer fixing: less) (auto intro: finite_linorder_min_induct)
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   886
qed
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   887
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   888
end
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   889
d516da3e7c42 material from $AFP/Formula_Derivatives/FSet_More
Lars Hupel <lars.hupel@mytum.de>
parents: 66262
diff changeset
   890
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60228
diff changeset
   891
subsection \<open>Choice in fsets\<close>
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   892
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
   893
lemma fset_choice:
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   894
  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
   895
  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
   896
  using assms by transfer metis
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   897
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
   898
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60228
diff changeset
   899
subsection \<open>Induction and Cases rules for fsets\<close>
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   900
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   901
lemma fset_exhaust [case_names empty insert, cases type: fset]:
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
   902
  assumes fempty_case: "S = {||} \<Longrightarrow> P"
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   903
  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
   904
  shows "P"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   905
  using assms by transfer blast
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   906
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   907
lemma fset_induct [case_names empty insert]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   908
  assumes fempty_case: "P {||}"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   909
  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
   910
  shows "P S"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   911
proof -
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   912
  (* FIXME transfer and right_total vs. bi_total *)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   913
  note Domainp_forall_transfer[transfer_rule]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   914
  show ?thesis
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   915
  using assms by transfer (auto intro: finite_induct)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   916
qed
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   917
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   918
lemma fset_induct_stronger [case_names empty insert, induct type: fset]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   919
  assumes empty_fset_case: "P {||}"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   920
  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
   921
  shows "P S"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   922
proof -
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   923
  (* FIXME transfer and right_total vs. bi_total *)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   924
  note Domainp_forall_transfer[transfer_rule]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   925
  show ?thesis
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   926
  using assms by transfer (auto intro: finite_induct)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   927
qed
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 fset_card_induct:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   930
  assumes empty_fset_case: "P {||}"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   931
  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
   932
  shows "P S"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   933
proof (induct S)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   934
  case empty
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   935
  show "P {||}" by (rule empty_fset_case)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   936
next
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   937
  case (insert x S)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   938
  have h: "P S" by fact
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   939
  have "x |\<notin>| S" by fact
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
   940
  then have "Suc (fcard S) = fcard (finsert x S)"
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   941
    by transfer auto
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
   942
  then show "P (finsert x S)"
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   943
    using h card_fset_Suc_case by simp
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   944
qed
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   945
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   946
lemma fset_strong_cases:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   947
  obtains "xs = {||}"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   948
    | ys x where "x |\<notin>| ys" and "xs = finsert x ys"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   949
by transfer blast
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   950
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   951
lemma fset_induct2:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   952
  "P {||} {||} \<Longrightarrow>
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   953
  (\<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
   954
  (\<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
   955
  (\<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
   956
  P xsa ysa"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   957
  apply (induct xsa arbitrary: ysa)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   958
  apply (induct_tac x rule: fset_induct_stronger)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   959
  apply simp_all
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   960
  apply (induct_tac xa rule: fset_induct_stronger)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   961
  apply simp_all
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   962
  done
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   963
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
   964
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60228
diff changeset
   965
subsection \<open>Setup for Lifting/Transfer\<close>
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   966
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60228
diff changeset
   967
subsubsection \<open>Relator and predicator properties\<close>
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   968
55938
f20d1db5aa3c renamed 'set_rel' to 'rel_set'
blanchet
parents: 55933
diff changeset
   969
lift_definition rel_fset :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a fset \<Rightarrow> 'b fset \<Rightarrow> bool" is rel_set
f20d1db5aa3c renamed 'set_rel' to 'rel_set'
blanchet
parents: 55933
diff changeset
   970
parametric rel_set_transfer .
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   971
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
   972
lemma rel_fset_alt_def: "rel_fset R = (\<lambda>A B. (\<forall>x.\<exists>y. x|\<in>|A \<longrightarrow> y|\<in>|B \<and> R x y)
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   973
  \<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
   974
apply (rule ext)+
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   975
apply transfer'
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
   976
apply (subst rel_set_def[unfolded fun_eq_iff])
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   977
by blast
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   978
55938
f20d1db5aa3c renamed 'set_rel' to 'rel_set'
blanchet
parents: 55933
diff changeset
   979
lemma finite_rel_set:
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   980
  assumes fin: "finite X" "finite Z"
55938
f20d1db5aa3c renamed 'set_rel' to 'rel_set'
blanchet
parents: 55933
diff changeset
   981
  assumes R_S: "rel_set (R OO S) X Z"
f20d1db5aa3c renamed 'set_rel' to 'rel_set'
blanchet
parents: 55933
diff changeset
   982
  shows "\<exists>Y. finite Y \<and> rel_set R X Y \<and> rel_set S Y Z"
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   983
proof -
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   984
  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
   985
  apply atomize_elim
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   986
  apply (subst bchoice_iff[symmetric])
55938
f20d1db5aa3c renamed 'set_rel' to 'rel_set'
blanchet
parents: 55933
diff changeset
   987
  using R_S[unfolded rel_set_def OO_def] by blast
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
   988
56646
360a05d60761 added 'size' of finite sets
blanchet
parents: 56525
diff changeset
   989
  obtain g where g: "\<forall>z\<in>Z. S (g z) z \<and> (\<exists>x\<in>X. R x (g z))"
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   990
  apply atomize_elim
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   991
  apply (subst bchoice_iff[symmetric])
55938
f20d1db5aa3c renamed 'set_rel' to 'rel_set'
blanchet
parents: 55933
diff changeset
   992
  using R_S[unfolded rel_set_def OO_def] by blast
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
   993
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   994
  let ?Y = "f ` X \<union> g ` Z"
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   995
  have "finite ?Y" by (simp add: fin)
55938
f20d1db5aa3c renamed 'set_rel' to 'rel_set'
blanchet
parents: 55933
diff changeset
   996
  moreover have "rel_set R X ?Y"
f20d1db5aa3c renamed 'set_rel' to 'rel_set'
blanchet
parents: 55933
diff changeset
   997
    unfolding rel_set_def
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
   998
    using f g by clarsimp blast
55938
f20d1db5aa3c renamed 'set_rel' to 'rel_set'
blanchet
parents: 55933
diff changeset
   999
  moreover have "rel_set S ?Y Z"
f20d1db5aa3c renamed 'set_rel' to 'rel_set'
blanchet
parents: 55933
diff changeset
  1000
    unfolding rel_set_def
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1001
    using f g by clarsimp blast
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1002
  ultimately show ?thesis by metis
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1003
qed
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1004
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60228
diff changeset
  1005
subsubsection \<open>Transfer rules for the Transfer package\<close>
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1006
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60228
diff changeset
  1007
text \<open>Unconditional transfer rules\<close>
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1008
63343
fb5d8a50c641 bundle lifting_syntax;
wenzelm
parents: 63331
diff changeset
  1009
context includes lifting_syntax
53963
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
  1010
begin
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
  1011
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1012
lemmas fempty_transfer [transfer_rule] = empty_transfer[Transfer.transferred]
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1013
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1014
lemma finsert_transfer [transfer_rule]:
55933
12ee2c407dad renamed 'fset_rel' to 'rel_fset'
blanchet
parents: 55738
diff changeset
  1015
  "(A ===> rel_fset A ===> rel_fset A) finsert finsert"
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55943
diff changeset
  1016
  unfolding rel_fun_def rel_fset_alt_def by blast
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1017
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1018
lemma funion_transfer [transfer_rule]:
55933
12ee2c407dad renamed 'fset_rel' to 'rel_fset'
blanchet
parents: 55738
diff changeset
  1019
  "(rel_fset A ===> rel_fset A ===> rel_fset A) funion funion"
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55943
diff changeset
  1020
  unfolding rel_fun_def rel_fset_alt_def by blast
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1021
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1022
lemma ffUnion_transfer [transfer_rule]:
55933
12ee2c407dad renamed 'fset_rel' to 'rel_fset'
blanchet
parents: 55738
diff changeset
  1023
  "(rel_fset (rel_fset A) ===> rel_fset A) ffUnion ffUnion"
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55943
diff changeset
  1024
  unfolding rel_fun_def rel_fset_alt_def by transfer (simp, fast)
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1025
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1026
lemma fimage_transfer [transfer_rule]:
55933
12ee2c407dad renamed 'fset_rel' to 'rel_fset'
blanchet
parents: 55738
diff changeset
  1027
  "((A ===> B) ===> rel_fset A ===> rel_fset B) fimage fimage"
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55943
diff changeset
  1028
  unfolding rel_fun_def rel_fset_alt_def by simp blast
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1029
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1030
lemma fBall_transfer [transfer_rule]:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67091
diff changeset
  1031
  "(rel_fset A ===> (A ===> (=)) ===> (=)) fBall fBall"
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55943
diff changeset
  1032
  unfolding rel_fset_alt_def rel_fun_def by blast
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1033
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1034
lemma fBex_transfer [transfer_rule]:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67091
diff changeset
  1035
  "(rel_fset A ===> (A ===> (=)) ===> (=)) fBex fBex"
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55943
diff changeset
  1036
  unfolding rel_fset_alt_def rel_fun_def by blast
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1037
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1038
(* FIXME transfer doesn't work here *)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1039
lemma fPow_transfer [transfer_rule]:
55933
12ee2c407dad renamed 'fset_rel' to 'rel_fset'
blanchet
parents: 55738
diff changeset
  1040
  "(rel_fset A ===> rel_fset (rel_fset A)) fPow fPow"
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55943
diff changeset
  1041
  unfolding rel_fun_def
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55943
diff changeset
  1042
  using Pow_transfer[unfolded rel_fun_def, rule_format, Transfer.transferred]
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1043
  by blast
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1044
55933
12ee2c407dad renamed 'fset_rel' to 'rel_fset'
blanchet
parents: 55738
diff changeset
  1045
lemma rel_fset_transfer [transfer_rule]:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67091
diff changeset
  1046
  "((A ===> B ===> (=)) ===> rel_fset A ===> rel_fset B ===> (=))
55933
12ee2c407dad renamed 'fset_rel' to 'rel_fset'
blanchet
parents: 55738
diff changeset
  1047
    rel_fset rel_fset"
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55943
diff changeset
  1048
  unfolding rel_fun_def
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55943
diff changeset
  1049
  using rel_set_transfer[unfolded rel_fun_def,rule_format, Transfer.transferred, where A = A and B = B]
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1050
  by simp
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1051
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1052
lemma bind_transfer [transfer_rule]:
55933
12ee2c407dad renamed 'fset_rel' to 'rel_fset'
blanchet
parents: 55738
diff changeset
  1053
  "(rel_fset A ===> (A ===> rel_fset B) ===> rel_fset B) fbind fbind"
63092
a949b2a5f51d eliminated use of empty "assms";
wenzelm
parents: 63040
diff changeset
  1054
  unfolding rel_fun_def
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55943
diff changeset
  1055
  using bind_transfer[unfolded rel_fun_def, rule_format, Transfer.transferred] by blast
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1056
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60228
diff changeset
  1057
text \<open>Rules requiring bi-unique, bi-total or right-total relations\<close>
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1058
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1059
lemma fmember_transfer [transfer_rule]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1060
  assumes "bi_unique A"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67091
diff changeset
  1061
  shows "(A ===> rel_fset A ===> (=)) (|\<in>|) (|\<in>|)"
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55943
diff changeset
  1062
  using assms unfolding rel_fun_def rel_fset_alt_def bi_unique_def by metis
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1063
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1064
lemma finter_transfer [transfer_rule]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1065
  assumes "bi_unique A"
55933
12ee2c407dad renamed 'fset_rel' to 'rel_fset'
blanchet
parents: 55738
diff changeset
  1066
  shows "(rel_fset A ===> rel_fset A ===> rel_fset A) finter finter"
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55943
diff changeset
  1067
  using assms unfolding rel_fun_def
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55943
diff changeset
  1068
  using inter_transfer[unfolded rel_fun_def, rule_format, Transfer.transferred] by blast
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1069
53963
51e81874b6f6 fold and lemmas about cardinality
kuncar
parents: 53953
diff changeset
  1070
lemma fminus_transfer [transfer_rule]:
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1071
  assumes "bi_unique A"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67091
diff changeset
  1072
  shows "(rel_fset A ===> rel_fset A ===> rel_fset A) (|-|) (|-|)"
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55943
diff changeset
  1073
  using assms unfolding rel_fun_def
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55943
diff changeset
  1074
  using Diff_transfer[unfolded rel_fun_def, rule_format, Transfer.transferred] by blast
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1075
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1076
lemma fsubset_transfer [transfer_rule]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1077
  assumes "bi_unique A"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67091
diff changeset
  1078
  shows "(rel_fset A ===> rel_fset A ===> (=)) (|\<subseteq>|) (|\<subseteq>|)"
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55943
diff changeset
  1079
  using assms unfolding rel_fun_def
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55943
diff changeset
  1080
  using subset_transfer[unfolded rel_fun_def, rule_format, Transfer.transferred] by blast
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1081
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1082
lemma fSup_transfer [transfer_rule]:
55938
f20d1db5aa3c renamed 'set_rel' to 'rel_set'
blanchet
parents: 55933
diff changeset
  1083
  "bi_unique A \<Longrightarrow> (rel_set (rel_fset A) ===> rel_fset A) Sup Sup"
63092
a949b2a5f51d eliminated use of empty "assms";
wenzelm
parents: 63040
diff changeset
  1084
  unfolding rel_fun_def
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1085
  apply clarify
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1086
  apply transfer'
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55943
diff changeset
  1087
  using Sup_fset_transfer[unfolded rel_fun_def] by blast
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1088
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1089
(* FIXME: add right_total_fInf_transfer *)
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1090
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1091
lemma fInf_transfer [transfer_rule]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1092
  assumes "bi_unique A" and "bi_total A"
55938
f20d1db5aa3c renamed 'set_rel' to 'rel_set'
blanchet
parents: 55933
diff changeset
  1093
  shows "(rel_set (rel_fset A) ===> rel_fset A) Inf Inf"
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55943
diff changeset
  1094
  using assms unfolding rel_fun_def
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1095
  apply clarify
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1096
  apply transfer'
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55943
diff changeset
  1097
  using Inf_fset_transfer[unfolded rel_fun_def] by blast
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1098
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1099
lemma ffilter_transfer [transfer_rule]:
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1100
  assumes "bi_unique A"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67091
diff changeset
  1101
  shows "((A ===> (=)) ===> rel_fset A ===> rel_fset A) ffilter ffilter"
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55943
diff changeset
  1102
  using assms unfolding rel_fun_def
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55943
diff changeset
  1103
  using Lifting_Set.filter_transfer[unfolded rel_fun_def, rule_format, Transfer.transferred] by blast
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1104
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1105
lemma card_transfer [transfer_rule]:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67091
diff changeset
  1106
  "bi_unique A \<Longrightarrow> (rel_fset A ===> (=)) fcard fcard"
63092
a949b2a5f51d eliminated use of empty "assms";
wenzelm
parents: 63040
diff changeset
  1107
  unfolding rel_fun_def
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55943
diff changeset
  1108
  using card_transfer[unfolded rel_fun_def, rule_format, Transfer.transferred] by blast
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1109
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1110
end
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1111
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1112
lifting_update fset.lifting
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1113
lifting_forget fset.lifting
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1114
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1115
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60228
diff changeset
  1116
subsection \<open>BNF setup\<close>
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1117
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1118
context
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1119
includes fset.lifting
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1120
begin
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1121
55933
12ee2c407dad renamed 'fset_rel' to 'rel_fset'
blanchet
parents: 55738
diff changeset
  1122
lemma rel_fset_alt:
12ee2c407dad renamed 'fset_rel' to 'rel_fset'
blanchet
parents: 55738
diff changeset
  1123
  "rel_fset R a b \<longleftrightarrow> (\<forall>t \<in> fset a. \<exists>u \<in> fset b. R t u) \<and> (\<forall>t \<in> fset b. \<exists>u \<in> fset a. R u t)"
55938
f20d1db5aa3c renamed 'set_rel' to 'rel_set'
blanchet
parents: 55933
diff changeset
  1124
by transfer (simp add: rel_set_def)
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1125
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1126
lemma fset_to_fset: "finite A \<Longrightarrow> fset (the_inv fset A) = A"
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1127
apply (rule f_the_inv_into_f[unfolded inj_on_def])
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1128
apply (simp add: fset_inject)
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1129
apply (rule range_eqI Abs_fset_inverse[symmetric] CollectI)+
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1130
.
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1131
55933
12ee2c407dad renamed 'fset_rel' to 'rel_fset'
blanchet
parents: 55738
diff changeset
  1132
lemma rel_fset_aux:
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1133
"(\<forall>t \<in> fset a. \<exists>u \<in> fset b. R t u) \<and> (\<forall>u \<in> fset b. \<exists>t \<in> fset a. R t u) \<longleftrightarrow>
57398
882091eb1e9a merged two small theory files
blanchet
parents: 56651
diff changeset
  1134
 ((BNF_Def.Grp {a. fset a \<subseteq> {(a, b). R a b}} (fimage fst))\<inverse>\<inverse> OO
882091eb1e9a merged two small theory files
blanchet
parents: 56651
diff changeset
  1135
  BNF_Def.Grp {a. fset a \<subseteq> {(a, b). R a b}} (fimage snd)) a b" (is "?L = ?R")
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1136
proof
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1137
  assume ?L
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62390
diff changeset
  1138
  define R' where "R' =
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62390
diff changeset
  1139
    the_inv fset (Collect (case_prod R) \<inter> (fset a \<times> fset b))" (is "_ = the_inv fset ?L'")
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1140
  have "finite ?L'" by (intro finite_Int[OF disjI2] finite_cartesian_product) (transfer, simp)+
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1141
  hence *: "fset R' = ?L'" unfolding R'_def by (intro fset_to_fset)
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1142
  show ?R unfolding Grp_def relcompp.simps conversep.simps
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55129
diff changeset
  1143
  proof (intro CollectI case_prodI exI[of _ a] exI[of _ b] exI[of _ R'] conjI refl)
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60228
diff changeset
  1144
    from * show "a = fimage fst R'" using conjunct1[OF \<open>?L\<close>]
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1145
      by (transfer, auto simp add: image_def Int_def split: prod.splits)
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60228
diff changeset
  1146
    from * show "b = fimage snd R'" using conjunct2[OF \<open>?L\<close>]
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1147
      by (transfer, auto simp add: image_def Int_def split: prod.splits)
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1148
  qed (auto simp add: *)
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1149
next
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1150
  assume ?R thus ?L unfolding Grp_def relcompp.simps conversep.simps
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1151
  apply (simp add: subset_eq Ball_def)
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1152
  apply (rule conjI)
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1153
  apply (transfer, clarsimp, metis snd_conv)
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1154
  by (transfer, clarsimp, metis fst_conv)
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1155
qed
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1156
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1157
bnf "'a fset"
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1158
  map: fimage
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
  1159
  sets: fset
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1160
  bd: natLeq
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1161
  wits: "{||}"
55933
12ee2c407dad renamed 'fset_rel' to 'rel_fset'
blanchet
parents: 55738
diff changeset
  1162
  rel: rel_fset
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1163
apply -
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1164
          apply transfer' apply simp
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1165
         apply transfer' apply force
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1166
        apply transfer apply force
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1167
       apply transfer' apply force
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1168
      apply (rule natLeq_card_order)
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1169
     apply (rule natLeq_cinfinite)
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1170
    apply transfer apply (metis ordLess_imp_ordLeq finite_iff_ordLess_natLeq)
55933
12ee2c407dad renamed 'fset_rel' to 'rel_fset'
blanchet
parents: 55738
diff changeset
  1171
   apply (fastforce simp: rel_fset_alt)
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62087
diff changeset
  1172
 apply (simp add: Grp_def relcompp.simps conversep.simps fun_eq_iff rel_fset_alt
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
  1173
   rel_fset_aux[unfolded OO_Grp_alt])
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1174
apply transfer apply simp
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1175
done
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1176
55938
f20d1db5aa3c renamed 'set_rel' to 'rel_set'
blanchet
parents: 55933
diff changeset
  1177
lemma rel_fset_fset: "rel_set \<chi> (fset A1) (fset A2) = rel_fset \<chi> A1 A2"
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1178
  by transfer (rule refl)
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1179
53953
2f103a894ebe new theory of finite sets as a subtype
kuncar
parents:
diff changeset
  1180
end
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1181
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1182
lemmas [simp] = fset.map_comp fset.map_id fset.set_map
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1183
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1184
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60228
diff changeset
  1185
subsection \<open>Size setup\<close>
56646
360a05d60761 added 'size' of finite sets
blanchet
parents: 56525
diff changeset
  1186
360a05d60761 added 'size' of finite sets
blanchet
parents: 56525
diff changeset
  1187
context includes fset.lifting begin
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63649
diff changeset
  1188
lift_definition size_fset :: "('a \<Rightarrow> nat) \<Rightarrow> 'a fset \<Rightarrow> nat" is "\<lambda>f. sum (Suc \<circ> f)" .
56646
360a05d60761 added 'size' of finite sets
blanchet
parents: 56525
diff changeset
  1189
end
360a05d60761 added 'size' of finite sets
blanchet
parents: 56525
diff changeset
  1190
360a05d60761 added 'size' of finite sets
blanchet
parents: 56525
diff changeset
  1191
instantiation fset :: (type) size begin
360a05d60761 added 'size' of finite sets
blanchet
parents: 56525
diff changeset
  1192
definition size_fset where
360a05d60761 added 'size' of finite sets
blanchet
parents: 56525
diff changeset
  1193
  size_fset_overloaded_def: "size_fset = FSet.size_fset (\<lambda>_. 0)"
360a05d60761 added 'size' of finite sets
blanchet
parents: 56525
diff changeset
  1194
instance ..
360a05d60761 added 'size' of finite sets
blanchet
parents: 56525
diff changeset
  1195
end
360a05d60761 added 'size' of finite sets
blanchet
parents: 56525
diff changeset
  1196
360a05d60761 added 'size' of finite sets
blanchet
parents: 56525
diff changeset
  1197
lemmas size_fset_simps[simp] =
360a05d60761 added 'size' of finite sets
blanchet
parents: 56525
diff changeset
  1198
  size_fset_def[THEN meta_eq_to_obj_eq, THEN fun_cong, THEN fun_cong,
360a05d60761 added 'size' of finite sets
blanchet
parents: 56525
diff changeset
  1199
    unfolded map_fun_def comp_def id_apply]
360a05d60761 added 'size' of finite sets
blanchet
parents: 56525
diff changeset
  1200
360a05d60761 added 'size' of finite sets
blanchet
parents: 56525
diff changeset
  1201
lemmas size_fset_overloaded_simps[simp] =
360a05d60761 added 'size' of finite sets
blanchet
parents: 56525
diff changeset
  1202
  size_fset_simps[of "\<lambda>_. 0", unfolded add_0_left add_0_right,
360a05d60761 added 'size' of finite sets
blanchet
parents: 56525
diff changeset
  1203
    folded size_fset_overloaded_def]
360a05d60761 added 'size' of finite sets
blanchet
parents: 56525
diff changeset
  1204
360a05d60761 added 'size' of finite sets
blanchet
parents: 56525
diff changeset
  1205
lemma fset_size_o_map: "inj f \<Longrightarrow> size_fset g \<circ> fimage f = size_fset (g \<circ> f)"
60228
32dd7adba5a4 tuned proof; forget the transfer rule for size_fset
kuncar
parents: 58881
diff changeset
  1206
  apply (subst fun_eq_iff)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63649
diff changeset
  1207
  including fset.lifting by transfer (auto intro: sum.reindex_cong subset_inj_on)
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
  1208
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60228
diff changeset
  1209
setup \<open>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69164
diff changeset
  1210
BNF_LFP_Size.register_size_global \<^type_name>\<open>fset\<close> \<^const_name>\<open>size_fset\<close>
62082
614ef6d7a6b6 nicer 'Spec_Rules' for size function
blanchet
parents: 61952
diff changeset
  1211
  @{thm size_fset_overloaded_def} @{thms size_fset_simps size_fset_overloaded_simps}
614ef6d7a6b6 nicer 'Spec_Rules' for size function
blanchet
parents: 61952
diff changeset
  1212
  @{thms fset_size_o_map}
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60228
diff changeset
  1213
\<close>
56646
360a05d60761 added 'size' of finite sets
blanchet
parents: 56525
diff changeset
  1214
60228
32dd7adba5a4 tuned proof; forget the transfer rule for size_fset
kuncar
parents: 58881
diff changeset
  1215
lifting_update fset.lifting
32dd7adba5a4 tuned proof; forget the transfer rule for size_fset
kuncar
parents: 58881
diff changeset
  1216
lifting_forget fset.lifting
56646
360a05d60761 added 'size' of finite sets
blanchet
parents: 56525
diff changeset
  1217
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60228
diff changeset
  1218
subsection \<open>Advanced relator customization\<close>
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1219
67408
4a4c14b24800 prefer formal comments;
wenzelm
parents: 67399
diff changeset
  1220
text \<open>Set vs. sum relators:\<close>
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1221
63331
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
  1222
lemma rel_set_rel_sum[simp]:
247eac9758dd move Conditional_Complete_Lattices to Main
hoelzl
parents: 63092
diff changeset
  1223
"rel_set (rel_sum \<chi> \<phi>) A1 A2 \<longleftrightarrow>
55938
f20d1db5aa3c renamed 'set_rel' to 'rel_set'
blanchet
parents: 55933
diff changeset
  1224
 rel_set \<chi> (Inl -` A1) (Inl -` A2) \<and> rel_set \<phi> (Inr -` A1) (Inr -` A2)"
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1225
(is "?L \<longleftrightarrow> ?Rl \<and> ?Rr")
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1226
proof safe
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1227
  assume L: "?L"
55938
f20d1db5aa3c renamed 'set_rel' to 'rel_set'
blanchet
parents: 55933
diff changeset
  1228
  show ?Rl unfolding rel_set_def Bex_def vimage_eq proof safe
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1229
    fix l1 assume "Inl l1 \<in> A1"
55943
5c2df04e97d1 renamed 'sum_rel' to 'rel_sum'
blanchet
parents: 55938
diff changeset
  1230
    then obtain a2 where a2: "a2 \<in> A2" and "rel_sum \<chi> \<phi> (Inl l1) a2"
55938
f20d1db5aa3c renamed 'set_rel' to 'rel_set'
blanchet
parents: 55933
diff changeset
  1231
    using L unfolding rel_set_def by auto
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1232
    then obtain l2 where "a2 = Inl l2 \<and> \<chi> l1 l2" by (cases a2, auto)
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1233
    thus "\<exists> l2. Inl l2 \<in> A2 \<and> \<chi> l1 l2" using a2 by auto
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1234
  next
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1235
    fix l2 assume "Inl l2 \<in> A2"
55943
5c2df04e97d1 renamed 'sum_rel' to 'rel_sum'
blanchet
parents: 55938
diff changeset
  1236
    then obtain a1 where a1: "a1 \<in> A1" and "rel_sum \<chi> \<phi> a1 (Inl l2)"
55938
f20d1db5aa3c renamed 'set_rel' to 'rel_set'
blanchet
parents: 55933
diff changeset
  1237
    using L unfolding rel_set_def by auto
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1238
    then obtain l1 where "a1 = Inl l1 \<and> \<chi> l1 l2" by (cases a1, auto)
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1239
    thus "\<exists> l1. Inl l1 \<in> A1 \<and> \<chi> l1 l2" using a1 by auto
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1240
  qed
55938
f20d1db5aa3c renamed 'set_rel' to 'rel_set'
blanchet
parents: 55933
diff changeset
  1241
  show ?Rr unfolding rel_set_def Bex_def vimage_eq proof safe
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1242
    fix r1 assume "Inr r1 \<in> A1"
55943
5c2df04e97d1 renamed 'sum_rel' to 'rel_sum'
blanchet
parents: 55938
diff changeset
  1243
    then obtain a2 where a2: "a2 \<in> A2" and "rel_sum \<chi> \<phi> (Inr r1) a2"
55938
f20d1db5aa3c renamed 'set_rel' to 'rel_set'
blanchet
parents: 55933
diff changeset
  1244
    using L unfolding rel_set_def by auto
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1245
    then obtain r2 where "a2 = Inr r2 \<and> \<phi> r1 r2" by (cases a2, auto)
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1246
    thus "\<exists> r2. Inr r2 \<in> A2 \<and> \<phi> r1 r2" using a2 by auto
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1247
  next
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1248
    fix r2 assume "Inr r2 \<in> A2"
55943
5c2df04e97d1 renamed 'sum_rel' to 'rel_sum'
blanchet
parents: 55938
diff changeset
  1249
    then obtain a1 where a1: "a1 \<in> A1" and "rel_sum \<chi> \<phi> a1 (Inr r2)"
55938
f20d1db5aa3c renamed 'set_rel' to 'rel_set'
blanchet
parents: 55933
diff changeset
  1250
    using L unfolding rel_set_def by auto
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1251
    then obtain r1 where "a1 = Inr r1 \<and> \<phi> r1 r2" by (cases a1, auto)
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1252
    thus "\<exists> r1. Inr r1 \<in> A1 \<and> \<phi> r1 r2" using a1 by auto
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1253
  qed
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1254
next
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1255
  assume Rl: "?Rl" and Rr: "?Rr"
55938
f20d1db5aa3c renamed 'set_rel' to 'rel_set'
blanchet
parents: 55933
diff changeset
  1256
  show ?L unfolding rel_set_def Bex_def vimage_eq proof safe
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1257
    fix a1 assume a1: "a1 \<in> A1"
55943
5c2df04e97d1 renamed 'sum_rel' to 'rel_sum'
blanchet
parents: 55938
diff changeset
  1258
    show "\<exists> a2. a2 \<in> A2 \<and> rel_sum \<chi> \<phi> a1 a2"
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1259
    proof(cases a1)
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1260
      case (Inl l1) then obtain l2 where "Inl l2 \<in> A2 \<and> \<chi> l1 l2"
55938
f20d1db5aa3c renamed 'set_rel' to 'rel_set'
blanchet
parents: 55933
diff changeset
  1261
      using Rl a1 unfolding rel_set_def by blast
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1262
      thus ?thesis unfolding Inl by auto
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1263
    next
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1264
      case (Inr r1) then obtain r2 where "Inr r2 \<in> A2 \<and> \<phi> r1 r2"
55938
f20d1db5aa3c renamed 'set_rel' to 'rel_set'
blanchet
parents: 55933
diff changeset
  1265
      using Rr a1 unfolding rel_set_def by blast
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1266
      thus ?thesis unfolding Inr by auto
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1267
    qed
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1268
  next
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1269
    fix a2 assume a2: "a2 \<in> A2"
55943
5c2df04e97d1 renamed 'sum_rel' to 'rel_sum'
blanchet
parents: 55938
diff changeset
  1270
    show "\<exists> a1. a1 \<in> A1 \<and> rel_sum \<chi> \<phi> a1 a2"
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1271
    proof(cases a2)
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1272
      case (Inl l2) then obtain l1 where "Inl l1 \<in> A1 \<and> \<chi> l1 l2"
55938
f20d1db5aa3c renamed 'set_rel' to 'rel_set'
blanchet
parents: 55933
diff changeset
  1273
      using Rl a2 unfolding rel_set_def by blast
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1274
      thus ?thesis unfolding Inl by auto
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1275
    next
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1276
      case (Inr r2) then obtain r1 where "Inr r1 \<in> A1 \<and> \<phi> r1 r2"
55938
f20d1db5aa3c renamed 'set_rel' to 'rel_set'
blanchet
parents: 55933
diff changeset
  1277
      using Rr a2 unfolding rel_set_def by blast
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1278
      thus ?thesis unfolding Inr by auto
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1279
    qed
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1280
  qed
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1281
qed
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1282
60712
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1283
66262
4a2c9d32e7aa finite sets are countable
Lars Hupel <lars.hupel@mytum.de>
parents: 66261
diff changeset
  1284
subsubsection \<open>Countability\<close>
4a2c9d32e7aa finite sets are countable
Lars Hupel <lars.hupel@mytum.de>
parents: 66261
diff changeset
  1285
4a2c9d32e7aa finite sets are countable
Lars Hupel <lars.hupel@mytum.de>
parents: 66261
diff changeset
  1286
lemma exists_fset_of_list: "\<exists>xs. fset_of_list xs = S"
4a2c9d32e7aa finite sets are countable
Lars Hupel <lars.hupel@mytum.de>
parents: 66261
diff changeset
  1287
including fset.lifting
4a2c9d32e7aa finite sets are countable
Lars Hupel <lars.hupel@mytum.de>
parents: 66261
diff changeset
  1288
by transfer (rule finite_list)
4a2c9d32e7aa finite sets are countable
Lars Hupel <lars.hupel@mytum.de>
parents: 66261
diff changeset
  1289
4a2c9d32e7aa finite sets are countable
Lars Hupel <lars.hupel@mytum.de>
parents: 66261
diff changeset
  1290
lemma fset_of_list_surj[simp, intro]: "surj fset_of_list"
4a2c9d32e7aa finite sets are countable
Lars Hupel <lars.hupel@mytum.de>
parents: 66261
diff changeset
  1291
proof -
4a2c9d32e7aa finite sets are countable
Lars Hupel <lars.hupel@mytum.de>
parents: 66261
diff changeset
  1292
  have "x \<in> range fset_of_list" for x :: "'a fset"
4a2c9d32e7aa finite sets are countable
Lars Hupel <lars.hupel@mytum.de>
parents: 66261
diff changeset
  1293
    unfolding image_iff
4a2c9d32e7aa finite sets are countable
Lars Hupel <lars.hupel@mytum.de>
parents: 66261
diff changeset
  1294
    using exists_fset_of_list by fastforce
4a2c9d32e7aa finite sets are countable
Lars Hupel <lars.hupel@mytum.de>
parents: 66261
diff changeset
  1295
  thus ?thesis by auto
4a2c9d32e7aa finite sets are countable
Lars Hupel <lars.hupel@mytum.de>
parents: 66261
diff changeset
  1296
qed
4a2c9d32e7aa finite sets are countable
Lars Hupel <lars.hupel@mytum.de>
parents: 66261
diff changeset
  1297
4a2c9d32e7aa finite sets are countable
Lars Hupel <lars.hupel@mytum.de>
parents: 66261
diff changeset
  1298
instance fset :: (countable) countable
4a2c9d32e7aa finite sets are countable
Lars Hupel <lars.hupel@mytum.de>
parents: 66261
diff changeset
  1299
proof
4a2c9d32e7aa finite sets are countable
Lars Hupel <lars.hupel@mytum.de>
parents: 66261
diff changeset
  1300
  obtain to_nat :: "'a list \<Rightarrow> nat" where "inj to_nat"
4a2c9d32e7aa finite sets are countable
Lars Hupel <lars.hupel@mytum.de>
parents: 66261
diff changeset
  1301
    by (metis ex_inj)
4a2c9d32e7aa finite sets are countable
Lars Hupel <lars.hupel@mytum.de>
parents: 66261
diff changeset
  1302
  moreover have "inj (inv fset_of_list)"
4a2c9d32e7aa finite sets are countable
Lars Hupel <lars.hupel@mytum.de>
parents: 66261
diff changeset
  1303
    using fset_of_list_surj by (rule surj_imp_inj_inv)
4a2c9d32e7aa finite sets are countable
Lars Hupel <lars.hupel@mytum.de>
parents: 66261
diff changeset
  1304
  ultimately have "inj (to_nat \<circ> inv fset_of_list)"
69700
7a92cbec7030 new material about summations and powers, along with some tweaks
paulson <lp15@cam.ac.uk>
parents: 69654
diff changeset
  1305
    by (rule inj_compose)
66262
4a2c9d32e7aa finite sets are countable
Lars Hupel <lars.hupel@mytum.de>
parents: 66261
diff changeset
  1306
  thus "\<exists>to_nat::'a fset \<Rightarrow> nat. inj to_nat"
4a2c9d32e7aa finite sets are countable
Lars Hupel <lars.hupel@mytum.de>
parents: 66261
diff changeset
  1307
    by auto
4a2c9d32e7aa finite sets are countable
Lars Hupel <lars.hupel@mytum.de>
parents: 66261
diff changeset
  1308
qed
4a2c9d32e7aa finite sets are countable
Lars Hupel <lars.hupel@mytum.de>
parents: 66261
diff changeset
  1309
4a2c9d32e7aa finite sets are countable
Lars Hupel <lars.hupel@mytum.de>
parents: 66261
diff changeset
  1310
60712
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1311
subsection \<open>Quickcheck setup\<close>
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1312
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1313
text \<open>Setup adapted from sets.\<close>
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1314
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1315
notation Quickcheck_Exhaustive.orelse (infixr "orelse" 55)
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1316
72607
feebdaa346e5 bundles for reflected term syntax
haftmann
parents: 72581
diff changeset
  1317
context
feebdaa346e5 bundles for reflected term syntax
haftmann
parents: 72581
diff changeset
  1318
  includes term_syntax
feebdaa346e5 bundles for reflected term syntax
haftmann
parents: 72581
diff changeset
  1319
begin
feebdaa346e5 bundles for reflected term syntax
haftmann
parents: 72581
diff changeset
  1320
feebdaa346e5 bundles for reflected term syntax
haftmann
parents: 72581
diff changeset
  1321
definition [code_unfold]:
60712
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1322
"valterm_femptyset = Code_Evaluation.valtermify ({||} :: ('a :: typerep) fset)"
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1323
72607
feebdaa346e5 bundles for reflected term syntax
haftmann
parents: 72581
diff changeset
  1324
definition [code_unfold]:
60712
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1325
"valtermify_finsert x s = Code_Evaluation.valtermify finsert {\<cdot>} (x :: ('a :: typerep * _)) {\<cdot>} s"
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1326
72607
feebdaa346e5 bundles for reflected term syntax
haftmann
parents: 72581
diff changeset
  1327
end
feebdaa346e5 bundles for reflected term syntax
haftmann
parents: 72581
diff changeset
  1328
60712
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1329
instantiation fset :: (exhaustive) exhaustive
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1330
begin
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1331
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1332
fun exhaustive_fset where
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1333
"exhaustive_fset f i = (if i = 0 then None else (f {||} orelse exhaustive_fset (\<lambda>A. f A orelse Quickcheck_Exhaustive.exhaustive (\<lambda>x. if x |\<in>| A then None else f (finsert x A)) (i - 1)) (i - 1)))"
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1334
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1335
instance ..
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1336
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 54258
diff changeset
  1337
end
60712
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1338
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1339
instantiation fset :: (full_exhaustive) full_exhaustive
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1340
begin
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1341
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1342
fun full_exhaustive_fset where
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1343
"full_exhaustive_fset f i = (if i = 0 then None else (f valterm_femptyset orelse full_exhaustive_fset (\<lambda>A. f A orelse Quickcheck_Exhaustive.full_exhaustive (\<lambda>x. if fst x |\<in>| fst A then None else f (valtermify_finsert x A)) (i - 1)) (i - 1)))"
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1344
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1345
instance ..
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1346
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1347
end
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1348
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1349
no_notation Quickcheck_Exhaustive.orelse (infixr "orelse" 55)
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1350
72581
de581f98a3a1 bundled syntax for state monad combinators
haftmann
parents: 72302
diff changeset
  1351
instantiation fset :: (random) random
de581f98a3a1 bundled syntax for state monad combinators
haftmann
parents: 72302
diff changeset
  1352
begin
60712
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1353
72581
de581f98a3a1 bundled syntax for state monad combinators
haftmann
parents: 72302
diff changeset
  1354
context
de581f98a3a1 bundled syntax for state monad combinators
haftmann
parents: 72302
diff changeset
  1355
  includes state_combinator_syntax
60712
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1356
begin
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1357
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1358
fun random_aux_fset :: "natural \<Rightarrow> natural \<Rightarrow> natural \<times> natural \<Rightarrow> ('a fset \<times> (unit \<Rightarrow> term)) \<times> natural \<times> natural" where
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1359
"random_aux_fset 0 j = Quickcheck_Random.collapse (Random.select_weight [(1, Pair valterm_femptyset)])" |
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1360
"random_aux_fset (Code_Numeral.Suc i) j =
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1361
  Quickcheck_Random.collapse (Random.select_weight
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1362
    [(1, Pair valterm_femptyset),
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1363
     (Code_Numeral.Suc i,
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1364
      Quickcheck_Random.random j \<circ>\<rightarrow> (\<lambda>x. random_aux_fset i j \<circ>\<rightarrow> (\<lambda>s. Pair (valtermify_finsert x s))))])"
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1365
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1366
lemma [code]:
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1367
  "random_aux_fset i j =
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1368
    Quickcheck_Random.collapse (Random.select_weight [(1, Pair valterm_femptyset),
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1369
      (i, Quickcheck_Random.random j \<circ>\<rightarrow> (\<lambda>x. random_aux_fset (i - 1) j \<circ>\<rightarrow> (\<lambda>s. Pair (valtermify_finsert x s))))])"
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1370
proof (induct i rule: natural.induct)
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1371
  case zero
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1372
  show ?case by (subst select_weight_drop_zero[symmetric]) (simp add: less_natural_def)
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1373
next
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1374
  case (Suc i)
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1375
  show ?case by (simp only: random_aux_fset.simps Suc_natural_minus_one)
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1376
qed
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1377
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1378
definition "random_fset i = random_aux_fset i i"
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1379
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1380
instance ..
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1381
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1382
end
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1383
72581
de581f98a3a1 bundled syntax for state monad combinators
haftmann
parents: 72302
diff changeset
  1384
end
60712
3ba16d28449d Quickcheck setup for finite sets
Lars Hupel <lars.hupel@mytum.de>
parents: 60679
diff changeset
  1385
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67091
diff changeset
  1386
end