src/HOL/Quotient_Examples/FSet.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Fri, 15 Oct 2010 21:50:26 +0900
changeset 39996 c02078ff8691
parent 39995 849578dd6127
child 40030 9f8dcf6ef563
permissions -rw-r--r--
FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
36524
3909002beca5 Tuning the quotient examples
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36465
diff changeset
     1
(*  Title:      HOL/Quotient_Examples/FSet.thy
36465
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
     2
    Author:     Cezary Kaliszyk, TU Munich
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
     3
    Author:     Christian Urban, TU Munich
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     4
36465
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
     5
A reasoning infrastructure for the type of finite sets.
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     6
*)
36465
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
     7
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     8
theory FSet
36465
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
     9
imports Quotient_List
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    10
begin
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    11
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    12
text {* Definiton of List relation and the quotient type *}
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    13
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    14
fun
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    15
  list_eq :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infix "\<approx>" 50)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    16
where
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
    17
  "list_eq xs ys = (set xs = set ys)"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    18
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    19
lemma list_eq_equivp:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    20
  shows "equivp list_eq"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    21
  unfolding equivp_reflp_symp_transp
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    22
  unfolding reflp_def symp_def transp_def
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    23
  by auto
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    24
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    25
quotient_type
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    26
  'a fset = "'a list" / "list_eq"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    27
  by (rule list_eq_equivp)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    28
39994
7bd8013b903f FSet: give names to respectfulness theorems, rename list_all2_refl to avoid clash
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39302
diff changeset
    29
text {* Raw definitions of membership, sublist, cardinality,
7bd8013b903f FSet: give names to respectfulness theorems, rename list_all2_refl to avoid clash
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39302
diff changeset
    30
  intersection
7bd8013b903f FSet: give names to respectfulness theorems, rename list_all2_refl to avoid clash
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39302
diff changeset
    31
*}
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    32
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    33
definition
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    34
  memb :: "'a \<Rightarrow> 'a list \<Rightarrow> bool"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    35
where
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    36
  "memb x xs \<equiv> x \<in> set xs"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    37
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    38
definition
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    39
  sub_list :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    40
where
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
    41
  "sub_list xs ys \<equiv> set xs \<subseteq> set ys"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    42
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
    43
definition
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    44
  fcard_raw :: "'a list \<Rightarrow> nat"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    45
where
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
    46
  "fcard_raw xs = card (set xs)"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    47
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    48
primrec
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    49
  finter_raw :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    50
where
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
    51
  "finter_raw [] ys = []"
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
    52
| "finter_raw (x # xs) ys =
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
    53
    (if x \<in> set ys then x # (finter_raw xs ys) else finter_raw xs ys)"
36675
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
    54
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
    55
primrec
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
    56
  fminus_raw :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
    57
where
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
    58
  "fminus_raw ys [] = ys"
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
    59
| "fminus_raw ys (x # xs) = fminus_raw (removeAll x ys) xs"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    60
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    61
definition
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    62
  rsp_fold
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    63
where
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    64
  "rsp_fold f = (\<forall>u v w. (f u (f v w) = f v (f u w)))"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    65
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    66
primrec
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    67
  ffold_raw :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a list \<Rightarrow> 'b"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    68
where
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    69
  "ffold_raw f z [] = z"
36675
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
    70
| "ffold_raw f z (a # xs) =
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    71
     (if (rsp_fold f) then
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
    72
       if a \<in> set xs then ffold_raw f z xs
36675
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
    73
       else f a (ffold_raw f z xs)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    74
     else z)"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    75
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    76
text {* Composition Quotient *}
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    77
39994
7bd8013b903f FSet: give names to respectfulness theorems, rename list_all2_refl to avoid clash
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39302
diff changeset
    78
lemma list_all2_refl1:
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
    79
  shows "(list_all2 op \<approx>) r r"
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
    80
  by (rule list_all2_refl) (metis equivp_def fset_equivp)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    81
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    82
lemma compose_list_refl:
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
    83
  shows "(list_all2 op \<approx> OOO op \<approx>) r r"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    84
proof
36465
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
    85
  have *: "r \<approx> r" by (rule equivp_reflp[OF fset_equivp])
39994
7bd8013b903f FSet: give names to respectfulness theorems, rename list_all2_refl to avoid clash
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39302
diff changeset
    86
  show "list_all2 op \<approx> r r" by (rule list_all2_refl1)
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
    87
  with * show "(op \<approx> OO list_all2 op \<approx>) r r" ..
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    88
qed
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    89
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    90
lemma Quotient_fset_list:
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
    91
  shows "Quotient (list_all2 op \<approx>) (map abs_fset) (map rep_fset)"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    92
  by (fact list_quotient[OF Quotient_fset])
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    93
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    94
lemma map_rel_cong: "b \<approx> ba \<Longrightarrow> map f b \<approx> map f ba"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    95
  unfolding list_eq.simps
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
    96
  by (simp only: set_map)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    97
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    98
lemma quotient_compose_list[quot_thm]:
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
    99
  shows  "Quotient ((list_all2 op \<approx>) OOO (op \<approx>))
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   100
    (abs_fset \<circ> (map abs_fset)) ((map rep_fset) \<circ> rep_fset)"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   101
  unfolding Quotient_def comp_def
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   102
proof (intro conjI allI)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   103
  fix a r s
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   104
  show "abs_fset (map abs_fset (map rep_fset (rep_fset a))) = a"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   105
    by (simp add: abs_o_rep[OF Quotient_fset] Quotient_abs_rep[OF Quotient_fset] map_id)
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   106
  have b: "list_all2 op \<approx> (map rep_fset (rep_fset a)) (map rep_fset (rep_fset a))"
39994
7bd8013b903f FSet: give names to respectfulness theorems, rename list_all2_refl to avoid clash
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39302
diff changeset
   107
    by (rule list_all2_refl1)
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   108
  have c: "(op \<approx> OO list_all2 op \<approx>) (map rep_fset (rep_fset a)) (map rep_fset (rep_fset a))"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   109
    by (rule, rule equivp_reflp[OF fset_equivp]) (rule b)
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   110
  show "(list_all2 op \<approx> OOO op \<approx>) (map rep_fset (rep_fset a)) (map rep_fset (rep_fset a))"
39994
7bd8013b903f FSet: give names to respectfulness theorems, rename list_all2_refl to avoid clash
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39302
diff changeset
   111
    by (rule, rule list_all2_refl1) (rule c)
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   112
  show "(list_all2 op \<approx> OOO op \<approx>) r s = ((list_all2 op \<approx> OOO op \<approx>) r r \<and>
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   113
        (list_all2 op \<approx> OOO op \<approx>) s s \<and> abs_fset (map abs_fset r) = abs_fset (map abs_fset s))"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   114
  proof (intro iffI conjI)
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   115
    show "(list_all2 op \<approx> OOO op \<approx>) r r" by (rule compose_list_refl)
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   116
    show "(list_all2 op \<approx> OOO op \<approx>) s s" by (rule compose_list_refl)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   117
  next
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   118
    assume a: "(list_all2 op \<approx> OOO op \<approx>) r s"
36465
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   119
    then have b: "map abs_fset r \<approx> map abs_fset s"
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   120
    proof (elim pred_compE)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   121
      fix b ba
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   122
      assume c: "list_all2 op \<approx> r b"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   123
      assume d: "b \<approx> ba"
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   124
      assume e: "list_all2 op \<approx> ba s"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   125
      have f: "map abs_fset r = map abs_fset b"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   126
        using Quotient_rel[OF Quotient_fset_list] c by blast
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   127
      have "map abs_fset ba = map abs_fset s"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   128
        using Quotient_rel[OF Quotient_fset_list] e by blast
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   129
      then have g: "map abs_fset s = map abs_fset ba" by simp
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   130
      then show "map abs_fset r \<approx> map abs_fset s" using d f map_rel_cong by simp
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   131
    qed
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   132
    then show "abs_fset (map abs_fset r) = abs_fset (map abs_fset s)"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   133
      using Quotient_rel[OF Quotient_fset] by blast
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   134
  next
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   135
    assume a: "(list_all2 op \<approx> OOO op \<approx>) r r \<and> (list_all2 op \<approx> OOO op \<approx>) s s
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   136
      \<and> abs_fset (map abs_fset r) = abs_fset (map abs_fset s)"
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   137
    then have s: "(list_all2 op \<approx> OOO op \<approx>) s s" by simp
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   138
    have d: "map abs_fset r \<approx> map abs_fset s"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   139
      by (subst Quotient_rel[OF Quotient_fset]) (simp add: a)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   140
    have b: "map rep_fset (map abs_fset r) \<approx> map rep_fset (map abs_fset s)"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   141
      by (rule map_rel_cong[OF d])
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   142
    have y: "list_all2 op \<approx> (map rep_fset (map abs_fset s)) s"
39994
7bd8013b903f FSet: give names to respectfulness theorems, rename list_all2_refl to avoid clash
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39302
diff changeset
   143
      by (fact rep_abs_rsp_left[OF Quotient_fset_list, OF list_all2_refl1[of s]])
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   144
    have c: "(op \<approx> OO list_all2 op \<approx>) (map rep_fset (map abs_fset r)) s"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   145
      by (rule pred_compI) (rule b, rule y)
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   146
    have z: "list_all2 op \<approx> r (map rep_fset (map abs_fset r))"
39994
7bd8013b903f FSet: give names to respectfulness theorems, rename list_all2_refl to avoid clash
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39302
diff changeset
   147
      by (fact rep_abs_rsp[OF Quotient_fset_list, OF list_all2_refl1[of r]])
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   148
    then show "(list_all2 op \<approx> OOO op \<approx>) r s"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   149
      using a c pred_compI by simp
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   150
  qed
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   151
qed
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   152
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   153
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   154
lemma set_finter_raw[simp]:
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   155
  "set (finter_raw xs ys) = set xs \<inter> set ys"
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   156
  by (induct xs) (auto simp add: memb_def)
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   157
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   158
lemma set_fminus_raw[simp]: 
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   159
  "set (fminus_raw xs ys) = (set xs - set ys)"
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   160
  by (induct ys arbitrary: xs) (auto)
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   161
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   162
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   163
text {* Respectfullness *}
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   164
39994
7bd8013b903f FSet: give names to respectfulness theorems, rename list_all2_refl to avoid clash
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39302
diff changeset
   165
lemma append_rsp[quot_respect]:
7bd8013b903f FSet: give names to respectfulness theorems, rename list_all2_refl to avoid clash
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39302
diff changeset
   166
  shows "(op \<approx> ===> op \<approx> ===> op \<approx>) append append"
7bd8013b903f FSet: give names to respectfulness theorems, rename list_all2_refl to avoid clash
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39302
diff changeset
   167
  by (simp)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   168
39994
7bd8013b903f FSet: give names to respectfulness theorems, rename list_all2_refl to avoid clash
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39302
diff changeset
   169
lemma sub_list_rsp[quot_respect]:
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   170
  shows "(op \<approx> ===> op \<approx> ===> op =) sub_list sub_list"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   171
  by (auto simp add: sub_list_def)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   172
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   173
lemma memb_rsp[quot_respect]:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   174
  shows "(op = ===> op \<approx> ===> op =) memb memb"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   175
  by (auto simp add: memb_def)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   176
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   177
lemma nil_rsp[quot_respect]:
39994
7bd8013b903f FSet: give names to respectfulness theorems, rename list_all2_refl to avoid clash
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39302
diff changeset
   178
  shows "(op \<approx>) Nil Nil"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   179
  by simp
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   180
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   181
lemma cons_rsp[quot_respect]:
39994
7bd8013b903f FSet: give names to respectfulness theorems, rename list_all2_refl to avoid clash
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39302
diff changeset
   182
  shows "(op = ===> op \<approx> ===> op \<approx>) Cons Cons"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   183
  by simp
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   184
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   185
lemma map_rsp[quot_respect]:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   186
  shows "(op = ===> op \<approx> ===> op \<approx>) map map"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   187
  by auto
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   188
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   189
lemma set_rsp[quot_respect]:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   190
  "(op \<approx> ===> op =) set set"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   191
  by auto
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   192
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   193
lemma list_equiv_rsp[quot_respect]:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   194
  shows "(op \<approx> ===> op \<approx> ===> op =) op \<approx> op \<approx>"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   195
  by auto
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   196
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   197
lemma finter_raw_rsp[quot_respect]:
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   198
  shows "(op \<approx> ===> op \<approx> ===> op \<approx>) finter_raw finter_raw"
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   199
  by simp
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   200
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   201
lemma removeAll_rsp[quot_respect]:
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   202
  shows "(op = ===> op \<approx> ===> op \<approx>) removeAll removeAll"
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   203
  by simp
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   204
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   205
lemma fminus_raw_rsp[quot_respect]:
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   206
  shows "(op \<approx> ===> op \<approx> ===> op \<approx>) fminus_raw fminus_raw"
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   207
  by simp
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   208
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   209
lemma fcard_raw_rsp[quot_respect]:
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   210
  shows "(op \<approx> ===> op =) fcard_raw fcard_raw"
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   211
  by (simp add: fcard_raw_def)
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   212
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   213
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   214
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   215
lemma not_memb_nil:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   216
  shows "\<not> memb x []"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   217
  by (simp add: memb_def)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   218
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   219
lemma memb_cons_iff:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   220
  shows "memb x (y # xs) = (x = y \<or> memb x xs)"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   221
  by (induct xs) (auto simp add: memb_def)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   222
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   223
lemma memb_absorb:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   224
  shows "memb x xs \<Longrightarrow> x # xs \<approx> xs"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   225
  by (induct xs) (auto simp add: memb_def)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   226
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   227
lemma none_memb_nil:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   228
  "(\<forall>x. \<not> memb x xs) = (xs \<approx> [])"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   229
  by (simp add: memb_def)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   230
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   231
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   232
lemma memb_commute_ffold_raw:
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   233
  "rsp_fold f \<Longrightarrow> h \<in> set b \<Longrightarrow> ffold_raw f z b = f h (ffold_raw f z (removeAll h b))"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   234
  apply (induct b)
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   235
  apply (auto simp add: rsp_fold_def)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   236
  done
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   237
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   238
lemma ffold_raw_rsp_pre:
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   239
  "set a = set b \<Longrightarrow> ffold_raw f z a = ffold_raw f z b"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   240
  apply (induct a arbitrary: b)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   241
  apply (simp)
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   242
  apply (simp (no_asm_use))
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   243
  apply (rule conjI)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   244
  apply (rule_tac [!] impI)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   245
  apply (rule_tac [!] conjI)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   246
  apply (rule_tac [!] impI)
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   247
  apply (metis insert_absorb)
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   248
  apply (metis List.insert_def List.set.simps(2) List.set_insert ffold_raw.simps(2))
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   249
  apply (metis Diff_insert_absorb insertI1 memb_commute_ffold_raw set_removeAll)
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   250
  apply(drule_tac x="removeAll a1 b" in meta_spec)
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   251
  apply(auto)
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   252
  apply(drule meta_mp)
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   253
  apply(blast)
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   254
  by (metis List.set.simps(2) emptyE ffold_raw.simps(2) in_listsp_conv_set listsp.simps mem_def)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   255
39995
849578dd6127 FSet tuned
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39994
diff changeset
   256
lemma ffold_raw_rsp[quot_respect]:
849578dd6127 FSet tuned
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39994
diff changeset
   257
  shows "(op = ===> op = ===> op \<approx> ===> op =) ffold_raw ffold_raw"
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   258
  unfolding fun_rel_def
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   259
  by(auto intro: ffold_raw_rsp_pre)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   260
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   261
lemma concat_rsp_pre:
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   262
  assumes a: "list_all2 op \<approx> x x'"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   263
  and     b: "x' \<approx> y'"
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   264
  and     c: "list_all2 op \<approx> y' y"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   265
  and     d: "\<exists>x\<in>set x. xa \<in> set x"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   266
  shows "\<exists>x\<in>set y. xa \<in> set x"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   267
proof -
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   268
  obtain xb where e: "xb \<in> set x" and f: "xa \<in> set xb" using d by auto
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   269
  have "\<exists>y. y \<in> set x' \<and> xb \<approx> y" by (rule list_all2_find_element[OF e a])
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   270
  then obtain ya where h: "ya \<in> set x'" and i: "xb \<approx> ya" by auto
36465
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   271
  have "ya \<in> set y'" using b h by simp
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   272
  then have "\<exists>yb. yb \<in> set y \<and> ya \<approx> yb" using c by (rule list_all2_find_element)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   273
  then show ?thesis using f i by auto
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   274
qed
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   275
39994
7bd8013b903f FSet: give names to respectfulness theorems, rename list_all2_refl to avoid clash
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39302
diff changeset
   276
lemma concat_rsp[quot_respect]:
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   277
  shows "(list_all2 op \<approx> OOO op \<approx> ===> op \<approx>) concat concat"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   278
proof (rule fun_relI, elim pred_compE)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   279
  fix a b ba bb
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   280
  assume a: "list_all2 op \<approx> a ba"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   281
  assume b: "ba \<approx> bb"
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   282
  assume c: "list_all2 op \<approx> bb b"
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   283
  have "\<forall>x. (\<exists>xa\<in>set a. x \<in> set xa) = (\<exists>xa\<in>set b. x \<in> set xa)" 
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   284
  proof
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   285
    fix x
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   286
    show "(\<exists>xa\<in>set a. x \<in> set xa) = (\<exists>xa\<in>set b. x \<in> set xa)" 
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   287
    proof
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   288
      assume d: "\<exists>xa\<in>set a. x \<in> set xa"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   289
      show "\<exists>xa\<in>set b. x \<in> set xa" by (rule concat_rsp_pre[OF a b c d])
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   290
    next
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   291
      assume e: "\<exists>xa\<in>set b. x \<in> set xa"
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   292
      have a': "list_all2 op \<approx> ba a" by (rule list_all2_symp[OF list_eq_equivp, OF a])
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   293
      have b': "bb \<approx> ba" by (rule equivp_symp[OF list_eq_equivp, OF b])
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   294
      have c': "list_all2 op \<approx> b bb" by (rule list_all2_symp[OF list_eq_equivp, OF c])
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   295
      show "\<exists>xa\<in>set a. x \<in> set xa" by (rule concat_rsp_pre[OF c' b' a' e])
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   296
    qed
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   297
  qed
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   298
  then show "concat a \<approx> concat b" by auto
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   299
qed
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   300
36639
6243b49498ea added function ffilter and some lemmas from Finite_Set to the FSet theory
bulwahn
parents: 36524
diff changeset
   301
lemma [quot_respect]:
39994
7bd8013b903f FSet: give names to respectfulness theorems, rename list_all2_refl to avoid clash
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39302
diff changeset
   302
  shows "((op =) ===> op \<approx> ===> op \<approx>) filter filter"
36639
6243b49498ea added function ffilter and some lemmas from Finite_Set to the FSet theory
bulwahn
parents: 36524
diff changeset
   303
  by auto
6243b49498ea added function ffilter and some lemmas from Finite_Set to the FSet theory
bulwahn
parents: 36524
diff changeset
   304
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   305
text {* Distributive lattice with bot *}
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   306
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   307
lemma append_inter_distrib:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   308
  "x @ (finter_raw y z) \<approx> finter_raw (x @ y) (x @ z)"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   309
  apply (induct x)
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   310
  apply (auto)
36465
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   311
  done
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   312
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   313
instantiation fset :: (type) "{bounded_lattice_bot, distrib_lattice, minus}"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   314
begin
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   315
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   316
quotient_definition
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   317
  "bot :: 'a fset" is "[] :: 'a list"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   318
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   319
abbreviation
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   320
  fempty  ("{||}")
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   321
where
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   322
  "{||} \<equiv> bot :: 'a fset"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   323
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   324
quotient_definition
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   325
  "less_eq_fset \<Colon> ('a fset \<Rightarrow> 'a fset \<Rightarrow> bool)"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   326
is
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   327
  "sub_list \<Colon> ('a list \<Rightarrow> 'a list \<Rightarrow> bool)"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   328
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   329
abbreviation
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   330
  f_subset_eq :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<subseteq>|" 50)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   331
where
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   332
  "xs |\<subseteq>| ys \<equiv> xs \<le> ys"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   333
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   334
definition
39995
849578dd6127 FSet tuned
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39994
diff changeset
   335
  less_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool"
849578dd6127 FSet tuned
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39994
diff changeset
   336
where  
849578dd6127 FSet tuned
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39994
diff changeset
   337
  "xs < ys \<equiv> xs \<le> ys \<and> xs \<noteq> (ys::'a fset)"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   338
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   339
abbreviation
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   340
  fsubset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<subset>|" 50)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   341
where
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   342
  "xs |\<subset>| ys \<equiv> xs < ys"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   343
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   344
quotient_definition
39995
849578dd6127 FSet tuned
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39994
diff changeset
   345
  "sup :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   346
is
39994
7bd8013b903f FSet: give names to respectfulness theorems, rename list_all2_refl to avoid clash
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39302
diff changeset
   347
  "append :: 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   348
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   349
abbreviation
36675
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
   350
  funion (infixl "|\<union>|" 65)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   351
where
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   352
  "xs |\<union>| ys \<equiv> sup (xs :: 'a fset) ys"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   353
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   354
quotient_definition
39995
849578dd6127 FSet tuned
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39994
diff changeset
   355
  "inf :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   356
is
39995
849578dd6127 FSet tuned
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39994
diff changeset
   357
  "finter_raw :: 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   358
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   359
abbreviation
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   360
  finter (infixl "|\<inter>|" 65)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   361
where
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   362
  "xs |\<inter>| ys \<equiv> inf (xs :: 'a fset) ys"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   363
36675
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
   364
quotient_definition
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   365
  "minus :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
36675
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
   366
is
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   367
  "fminus_raw :: 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
36675
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
   368
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   369
instance
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   370
proof
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   371
  fix x y z :: "'a fset"
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   372
  show "x |\<subset>| y \<longleftrightarrow> x |\<subseteq>| y \<and> \<not> y |\<subseteq>| x"
39995
849578dd6127 FSet tuned
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39994
diff changeset
   373
    unfolding less_fset_def 
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   374
    by (descending) (auto simp add: sub_list_def)
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   375
  show "x |\<subseteq>| x"  by (descending) (simp add: sub_list_def)
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   376
  show "{||} |\<subseteq>| x" by (descending) (simp add: sub_list_def)
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   377
  show "x |\<subseteq>| x |\<union>| y" by (descending) (simp add: sub_list_def)
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   378
  show "y |\<subseteq>| x |\<union>| y" by (descending) (simp add: sub_list_def)
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   379
  show "x |\<inter>| y |\<subseteq>| x"
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   380
    by (descending) (simp add: sub_list_def memb_def[symmetric])
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   381
  show "x |\<inter>| y |\<subseteq>| y" 
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   382
    by (descending) (simp add: sub_list_def memb_def[symmetric])
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   383
  show "x |\<union>| (y |\<inter>| z) = x |\<union>| y |\<inter>| (x |\<union>| z)" 
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   384
    by (descending) (rule append_inter_distrib)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   385
next
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   386
  fix x y z :: "'a fset"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   387
  assume a: "x |\<subseteq>| y"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   388
  assume b: "y |\<subseteq>| z"
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   389
  show "x |\<subseteq>| z" using a b 
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   390
    by (descending) (simp add: sub_list_def)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   391
next
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   392
  fix x y :: "'a fset"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   393
  assume a: "x |\<subseteq>| y"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   394
  assume b: "y |\<subseteq>| x"
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   395
  show "x = y" using a b 
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   396
    by (descending) (unfold sub_list_def list_eq.simps, blast)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   397
next
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   398
  fix x y z :: "'a fset"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   399
  assume a: "y |\<subseteq>| x"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   400
  assume b: "z |\<subseteq>| x"
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   401
  show "y |\<union>| z |\<subseteq>| x" using a b 
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   402
    by (descending) (simp add: sub_list_def)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   403
next
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   404
  fix x y z :: "'a fset"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   405
  assume a: "x |\<subseteq>| y"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   406
  assume b: "x |\<subseteq>| z"
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   407
  show "x |\<subseteq>| y |\<inter>| z" using a b 
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   408
    by (descending) (simp add: sub_list_def memb_def[symmetric])
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   409
qed
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   410
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   411
end
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   412
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   413
section {* Finsert and Membership *}
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   414
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   415
quotient_definition
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   416
  "finsert :: 'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
39994
7bd8013b903f FSet: give names to respectfulness theorems, rename list_all2_refl to avoid clash
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39302
diff changeset
   417
is "Cons"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   418
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   419
syntax
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   420
  "@Finset"     :: "args => 'a fset"  ("{|(_)|}")
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   421
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   422
translations
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   423
  "{|x, xs|}" == "CONST finsert x {|xs|}"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   424
  "{|x|}"     == "CONST finsert x {||}"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   425
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   426
quotient_definition
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   427
  fin (infix "|\<in>|" 50)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   428
where
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   429
  "fin :: 'a \<Rightarrow> 'a fset \<Rightarrow> bool" is "memb"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   430
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   431
abbreviation
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   432
  fnotin :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<notin>|" 50)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   433
where
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   434
  "x |\<notin>| S \<equiv> \<not> (x |\<in>| S)"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   435
36465
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   436
section {* Other constants on the Quotient Type *}
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   437
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   438
quotient_definition
36465
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   439
  "fcard :: 'a fset \<Rightarrow> nat"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   440
is
39995
849578dd6127 FSet tuned
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39994
diff changeset
   441
  fcard_raw
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   442
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   443
quotient_definition
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   444
  "fmap :: ('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   445
is
39995
849578dd6127 FSet tuned
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39994
diff changeset
   446
  map
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   447
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   448
quotient_definition
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   449
  "fdelete :: 'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   450
  is removeAll
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   451
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   452
quotient_definition
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   453
  "fset :: 'a fset \<Rightarrow> 'a set"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   454
  is "set"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   455
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   456
quotient_definition
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   457
  "ffold :: ('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a fset \<Rightarrow> 'b"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   458
  is "ffold_raw"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   459
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   460
quotient_definition
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   461
  "fconcat :: ('a fset) fset \<Rightarrow> 'a fset"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   462
is
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   463
  "concat"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   464
36639
6243b49498ea added function ffilter and some lemmas from Finite_Set to the FSet theory
bulwahn
parents: 36524
diff changeset
   465
quotient_definition
6243b49498ea added function ffilter and some lemmas from Finite_Set to the FSet theory
bulwahn
parents: 36524
diff changeset
   466
  "ffilter :: ('a \<Rightarrow> bool) \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
6243b49498ea added function ffilter and some lemmas from Finite_Set to the FSet theory
bulwahn
parents: 36524
diff changeset
   467
is
6243b49498ea added function ffilter and some lemmas from Finite_Set to the FSet theory
bulwahn
parents: 36524
diff changeset
   468
  "filter"
6243b49498ea added function ffilter and some lemmas from Finite_Set to the FSet theory
bulwahn
parents: 36524
diff changeset
   469
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   470
text {* Compositional Respectfullness and Preservation *}
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   471
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   472
lemma [quot_respect]: "(list_all2 op \<approx> OOO op \<approx>) [] []"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   473
  by (fact compose_list_refl)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   474
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   475
lemma [quot_preserve]: "(abs_fset \<circ> map f) [] = abs_fset []"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   476
  by simp
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   477
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   478
lemma [quot_respect]:
39994
7bd8013b903f FSet: give names to respectfulness theorems, rename list_all2_refl to avoid clash
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39302
diff changeset
   479
  shows "(op \<approx> ===> list_all2 op \<approx> OOO op \<approx> ===> list_all2 op \<approx> OOO op \<approx>) Cons Cons"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   480
  apply auto
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   481
  apply (rule_tac b="x # b" in pred_compI)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   482
  apply auto
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   483
  apply (rule_tac b="x # ba" in pred_compI)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   484
  apply auto
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   485
  done
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   486
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   487
lemma [quot_preserve]:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   488
  "(rep_fset ---> (map rep_fset \<circ> rep_fset) ---> (abs_fset \<circ> map abs_fset)) op # = finsert"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
   489
  by (simp add: fun_eq_iff Quotient_abs_rep[OF Quotient_fset]
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   490
      abs_o_rep[OF Quotient_fset] map_id finsert_def)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   491
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   492
lemma [quot_preserve]:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   493
  "((map rep_fset \<circ> rep_fset) ---> (map rep_fset \<circ> rep_fset) ---> (abs_fset \<circ> map abs_fset)) op @ = funion"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
   494
  by (simp add: fun_eq_iff Quotient_abs_rep[OF Quotient_fset]
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   495
      abs_o_rep[OF Quotient_fset] map_id sup_fset_def)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   496
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   497
lemma list_all2_app_l:
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   498
  assumes a: "reflp R"
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   499
  and b: "list_all2 R l r"
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   500
  shows "list_all2 R (z @ l) (z @ r)"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   501
  by (induct z) (simp_all add: b rev_iffD1[OF a meta_eq_to_obj_eq[OF reflp_def]])
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   502
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   503
lemma append_rsp2_pre0:
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   504
  assumes a:"list_all2 op \<approx> x x'"
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   505
  shows "list_all2 op \<approx> (x @ z) (x' @ z)"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   506
  using a apply (induct x x' rule: list_induct2')
39994
7bd8013b903f FSet: give names to respectfulness theorems, rename list_all2_refl to avoid clash
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39302
diff changeset
   507
  by simp_all (rule list_all2_refl1)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   508
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   509
lemma append_rsp2_pre1:
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   510
  assumes a:"list_all2 op \<approx> x x'"
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   511
  shows "list_all2 op \<approx> (z @ x) (z @ x')"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   512
  using a apply (induct x x' arbitrary: z rule: list_induct2')
39994
7bd8013b903f FSet: give names to respectfulness theorems, rename list_all2_refl to avoid clash
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39302
diff changeset
   513
  apply (rule list_all2_refl1)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   514
  apply (simp_all del: list_eq.simps)
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   515
  apply (rule list_all2_app_l)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   516
  apply (simp_all add: reflp_def)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   517
  done
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   518
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   519
lemma append_rsp2_pre:
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   520
  assumes a:"list_all2 op \<approx> x x'"
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   521
  and     b: "list_all2 op \<approx> z z'"
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   522
  shows "list_all2 op \<approx> (x @ z) (x' @ z')"
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   523
  apply (rule list_all2_transp[OF fset_equivp])
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   524
  apply (rule append_rsp2_pre0)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   525
  apply (rule a)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   526
  using b apply (induct z z' rule: list_induct2')
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   527
  apply (simp_all only: append_Nil2)
39994
7bd8013b903f FSet: give names to respectfulness theorems, rename list_all2_refl to avoid clash
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39302
diff changeset
   528
  apply (rule list_all2_refl1)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   529
  apply simp_all
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   530
  apply (rule append_rsp2_pre1)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   531
  apply simp
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   532
  done
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   533
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   534
lemma [quot_respect]:
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   535
  "(list_all2 op \<approx> OOO op \<approx> ===> list_all2 op \<approx> OOO op \<approx> ===> list_all2 op \<approx> OOO op \<approx>) op @ op @"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   536
proof (intro fun_relI, elim pred_compE)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   537
  fix x y z w x' z' y' w' :: "'a list list"
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   538
  assume a:"list_all2 op \<approx> x x'"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   539
  and b:    "x' \<approx> y'"
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   540
  and c:    "list_all2 op \<approx> y' y"
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   541
  assume aa: "list_all2 op \<approx> z z'"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   542
  and bb:   "z' \<approx> w'"
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   543
  and cc:   "list_all2 op \<approx> w' w"
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   544
  have a': "list_all2 op \<approx> (x @ z) (x' @ z')" using a aa append_rsp2_pre by auto
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   545
  have b': "x' @ z' \<approx> y' @ w'" using b bb by simp
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   546
  have c': "list_all2 op \<approx> (y' @ w') (y @ w)" using c cc append_rsp2_pre by auto
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   547
  have d': "(op \<approx> OO list_all2 op \<approx>) (x' @ z') (y @ w)"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   548
    by (rule pred_compI) (rule b', rule c')
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36675
diff changeset
   549
  show "(list_all2 op \<approx> OOO op \<approx>) (x @ z) (y @ w)"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   550
    by (rule pred_compI) (rule a', rule d')
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   551
qed
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   552
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   553
text {* Raw theorems. Finsert, memb, singleron, sub_list *}
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   554
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   555
lemma nil_not_cons:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   556
  shows "\<not> ([] \<approx> x # xs)"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   557
  and   "\<not> (x # xs \<approx> [])"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   558
  by auto
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   559
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   560
lemma no_memb_nil:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   561
  "(\<forall>x. \<not> memb x xs) = (xs = [])"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   562
  by (simp add: memb_def)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   563
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   564
lemma memb_consI1:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   565
  shows "memb x (x # xs)"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   566
  by (simp add: memb_def)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   567
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   568
lemma memb_consI2:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   569
  shows "memb x xs \<Longrightarrow> memb x (y # xs)"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   570
  by (simp add: memb_def)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   571
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   572
lemma singleton_list_eq:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   573
  shows "[x] \<approx> [y] \<longleftrightarrow> x = y"
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   574
  by (simp)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   575
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   576
lemma sub_list_cons:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   577
  "sub_list (x # xs) ys = (memb x ys \<and> sub_list xs ys)"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   578
  by (auto simp add: memb_def sub_list_def)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   579
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   580
lemma fminus_raw_red: 
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   581
  "fminus_raw (x # xs) ys = (if x \<in> set ys then fminus_raw xs ys else x # (fminus_raw xs ys))"
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   582
  by (induct ys arbitrary: xs x) (simp_all)
36675
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
   583
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   584
text {* Cardinality of finite sets *}
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   585
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   586
lemma fcard_raw_0:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   587
  shows "fcard_raw xs = 0 \<longleftrightarrow> xs \<approx> []"
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   588
  unfolding fcard_raw_def
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   589
  by (induct xs) (auto)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   590
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   591
lemma memb_card_not_0:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   592
  assumes a: "memb a A"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   593
  shows "\<not>(fcard_raw A = 0)"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   594
proof -
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   595
  have "\<not>(\<forall>x. \<not> memb x A)" using a by auto
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   596
  then have "\<not>A \<approx> []" using none_memb_nil[of A] by simp
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   597
  then show ?thesis using fcard_raw_0[of A] by simp
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   598
qed
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   599
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   600
text {* fmap *}
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   601
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   602
lemma map_append:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   603
  "map f (xs @ ys) \<approx> (map f xs) @ (map f ys)"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   604
  by simp
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   605
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   606
lemma memb_append:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   607
  "memb x (xs @ ys) \<longleftrightarrow> memb x xs \<or> memb x ys"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   608
  by (induct xs) (simp_all add: not_memb_nil memb_cons_iff)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   609
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   610
lemma fset_raw_strong_cases:
36465
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   611
  obtains "xs = []"
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   612
    | x ys where "\<not> memb x ys" and "xs \<approx> x # ys"
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   613
proof (induct xs arbitrary: x ys)
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   614
  case Nil
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   615
  then show thesis by simp
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   616
next
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   617
  case (Cons a xs)
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   618
  have a: "\<lbrakk>xs = [] \<Longrightarrow> thesis; \<And>x ys. \<lbrakk>\<not> memb x ys; xs \<approx> x # ys\<rbrakk> \<Longrightarrow> thesis\<rbrakk> \<Longrightarrow> thesis" by fact
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   619
  have b: "\<And>x' ys'. \<lbrakk>\<not> memb x' ys'; a # xs \<approx> x' # ys'\<rbrakk> \<Longrightarrow> thesis" by fact
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   620
  have c: "xs = [] \<Longrightarrow> thesis" by (metis no_memb_nil singleton_list_eq b)
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   621
  have "\<And>x ys. \<lbrakk>\<not> memb x ys; xs \<approx> x # ys\<rbrakk> \<Longrightarrow> thesis"
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   622
  proof -
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   623
    fix x :: 'a
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   624
    fix ys :: "'a list"
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   625
    assume d:"\<not> memb x ys"
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   626
    assume e:"xs \<approx> x # ys"
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   627
    show thesis
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   628
    proof (cases "x = a")
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   629
      assume h: "x = a"
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   630
      then have f: "\<not> memb a ys" using d by simp
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   631
      have g: "a # xs \<approx> a # ys" using e h by auto
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   632
      show thesis using b f g by simp
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   633
    next
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   634
      assume h: "x \<noteq> a"
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   635
      then have f: "\<not> memb x (a # ys)" using d unfolding memb_def by auto
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   636
      have g: "a # xs \<approx> x # (a # ys)" using e h by auto
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   637
      show thesis using b f g by simp
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   638
    qed
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   639
  qed
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   640
  then show thesis using a c by blast
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   641
qed
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   642
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   643
section {* deletion *}
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   644
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   645
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   646
lemma fset_raw_removeAll_cases:
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   647
  "xs = [] \<or> (\<exists>x. memb x xs \<and> xs \<approx> x # removeAll x xs)"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   648
  by (induct xs) (auto simp add: memb_def)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   649
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   650
lemma fremoveAll_filter:
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   651
  "removeAll y xs = [x \<leftarrow> xs. x \<noteq> y]"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   652
  by (induct xs) simp_all
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   653
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   654
lemma fcard_raw_delete:
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   655
  "fcard_raw (removeAll y xs) = (if memb y xs then fcard_raw xs - 1 else fcard_raw xs)"
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   656
  by (auto simp add: fcard_raw_def memb_def)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   657
36465
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   658
lemma set_cong:
36646
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
   659
  shows "(x \<approx> y) = (set x = set y)"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   660
  by auto
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   661
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   662
lemma inj_map_eq_iff:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   663
  "inj f \<Longrightarrow> (map f l \<approx> map f m) = (l \<approx> m)"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
   664
  by (simp add: set_eq_iff[symmetric] inj_image_eq_iff)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   665
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   666
text {* alternate formulation with a different decomposition principle
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   667
  and a proof of equivalence *}
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   668
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   669
inductive
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   670
  list_eq2
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   671
where
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   672
  "list_eq2 (a # b # xs) (b # a # xs)"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   673
| "list_eq2 [] []"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   674
| "list_eq2 xs ys \<Longrightarrow> list_eq2 ys xs"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   675
| "list_eq2 (a # a # xs) (a # xs)"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   676
| "list_eq2 xs ys \<Longrightarrow> list_eq2 (a # xs) (a # ys)"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   677
| "\<lbrakk>list_eq2 xs1 xs2; list_eq2 xs2 xs3\<rbrakk> \<Longrightarrow> list_eq2 xs1 xs3"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   678
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   679
lemma list_eq2_refl:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   680
  shows "list_eq2 xs xs"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   681
  by (induct xs) (auto intro: list_eq2.intros)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   682
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   683
lemma cons_delete_list_eq2:
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   684
  shows "list_eq2 (a # (removeAll a A)) (if memb a A then A else a # A)"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   685
  apply (induct A)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   686
  apply (simp add: memb_def list_eq2_refl)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   687
  apply (case_tac "memb a (aa # A)")
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   688
  apply (simp_all only: memb_cons_iff)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   689
  apply (case_tac [!] "a = aa")
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   690
  apply (simp_all)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   691
  apply (case_tac "memb a A")
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   692
  apply (auto simp add: memb_def)[2]
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   693
  apply (metis list_eq2.intros(3) list_eq2.intros(4) list_eq2.intros(5) list_eq2.intros(6))
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   694
  apply (metis list_eq2.intros(1) list_eq2.intros(5) list_eq2.intros(6))
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   695
  apply (auto simp add: list_eq2_refl memb_def)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   696
  done
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   697
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   698
lemma memb_delete_list_eq2:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   699
  assumes a: "memb e r"
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   700
  shows "list_eq2 (e # removeAll e r) r"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   701
  using a cons_delete_list_eq2[of e r]
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   702
  by simp
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   703
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   704
lemma list_eq2_equiv:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   705
  "(l \<approx> r) \<longleftrightarrow> (list_eq2 l r)"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   706
proof
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   707
  show "list_eq2 l r \<Longrightarrow> l \<approx> r" by (induct rule: list_eq2.induct) auto
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   708
next
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   709
  {
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   710
    fix n
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   711
    assume a: "fcard_raw l = n" and b: "l \<approx> r"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   712
    have "list_eq2 l r"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   713
      using a b
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   714
    proof (induct n arbitrary: l r)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   715
      case 0
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   716
      have "fcard_raw l = 0" by fact
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   717
      then have "\<forall>x. \<not> memb x l" using memb_card_not_0[of _ l] by auto
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   718
      then have z: "l = []" using no_memb_nil by auto
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   719
      then have "r = []" using `l \<approx> r` by simp
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   720
      then show ?case using z list_eq2_refl by simp
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   721
    next
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   722
      case (Suc m)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   723
      have b: "l \<approx> r" by fact
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   724
      have d: "fcard_raw l = Suc m" by fact
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   725
      then have "\<exists>a. memb a l" 
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   726
	apply(simp add: fcard_raw_def memb_def)
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   727
	apply(drule card_eq_SucD)
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   728
	apply(blast)
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   729
	done
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   730
      then obtain a where e: "memb a l" by auto
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   731
      then have e': "memb a r" using list_eq.simps[simplified memb_def[symmetric], of l r] b 
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   732
	unfolding memb_def by auto
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   733
      have f: "fcard_raw (removeAll a l) = m" using fcard_raw_delete[of a l] e d by simp
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   734
      have g: "removeAll a l \<approx> removeAll a r" using removeAll_rsp b by simp
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   735
      have "list_eq2 (removeAll a l) (removeAll a r)" by (rule Suc.hyps[OF f g])
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   736
      then have h: "list_eq2 (a # removeAll a l) (a # removeAll a r)" by (rule list_eq2.intros(5))
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   737
      have i: "list_eq2 l (a # removeAll a l)"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   738
        by (rule list_eq2.intros(3)[OF memb_delete_list_eq2[OF e]])
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   739
      have "list_eq2 l (a # removeAll a r)" by (rule list_eq2.intros(6)[OF i h])
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   740
      then show ?case using list_eq2.intros(6)[OF _ memb_delete_list_eq2[OF e']] by simp
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   741
    qed
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   742
    }
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   743
  then show "l \<approx> r \<Longrightarrow> list_eq2 l r" by blast
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   744
qed
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   745
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   746
text {* Lifted theorems *}
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   747
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   748
lemma not_fin_fnil: "x |\<notin>| {||}"
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   749
  by (descending) (simp add: memb_def)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   750
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   751
lemma fin_finsert_iff[simp]:
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   752
  "x |\<in>| finsert y S \<longleftrightarrow> x = y \<or> x |\<in>| S"
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   753
  by (descending) (simp add: memb_def)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   754
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   755
lemma
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   756
  shows finsertI1: "x |\<in>| finsert x S"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   757
  and   finsertI2: "x |\<in>| S \<Longrightarrow> x |\<in>| finsert y S"
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   758
  by (lifting memb_consI1 memb_consI2)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   759
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   760
lemma finsert_absorb[simp]:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   761
  shows "x |\<in>| S \<Longrightarrow> finsert x S = S"
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   762
  by (descending) (auto simp add: memb_def)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   763
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   764
lemma fempty_not_finsert[simp]:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   765
  "{||} \<noteq> finsert x S"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   766
  "finsert x S \<noteq> {||}"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   767
  by (lifting nil_not_cons)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   768
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   769
lemma finsert_left_comm:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   770
  "finsert x (finsert y S) = finsert y (finsert x S)"
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   771
  by (descending) (auto)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   772
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   773
lemma finsert_left_idem:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   774
  "finsert x (finsert x S) = finsert x S"
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   775
  by (descending) (auto)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   776
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   777
lemma fsingleton_eq[simp]:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   778
  shows "{|x|} = {|y|} \<longleftrightarrow> x = y"
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   779
  by (descending) (auto)
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   780
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   781
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   782
text {* fset *}
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   783
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   784
lemma fset_simps[simp]:
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   785
  "fset {||} = ({} :: 'a set)"
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   786
  "fset (finsert (h :: 'a) t) = insert h (fset t)"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   787
  by (lifting set.simps)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   788
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   789
lemma in_fset:
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   790
  "x \<in> fset S \<equiv> x |\<in>| S"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   791
  by (lifting memb_def[symmetric])
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   792
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   793
lemma none_fin_fempty:
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   794
  "(\<forall>x. x |\<notin>| S) \<longleftrightarrow> S = {||}"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   795
  by (lifting none_memb_nil)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   796
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   797
lemma fset_cong:
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   798
  "S = T \<longleftrightarrow> fset S = fset T"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   799
  by (lifting set_cong)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   800
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   801
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   802
text {* fcard *}
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   803
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   804
lemma fcard_finsert_if [simp]:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   805
  shows "fcard (finsert x S) = (if x |\<in>| S then fcard S else Suc (fcard S))"
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   806
  by (descending) (auto simp add: fcard_raw_def memb_def insert_absorb)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   807
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   808
lemma fcard_0[simp]:
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   809
  shows "fcard S = 0 \<longleftrightarrow> S = {||}"
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   810
  by (descending) (simp add: fcard_raw_def)
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   811
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   812
lemma fcard_fempty[simp]:
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   813
  shows "fcard {||} = 0"
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   814
  by (simp add: fcard_0)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   815
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   816
lemma fcard_1:
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   817
  shows "fcard S = 1 \<longleftrightarrow> (\<exists>x. S = {|x|})"
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   818
  by (descending) (auto simp add: fcard_raw_def card_Suc_eq)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   819
36465
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   820
lemma fcard_gt_0:
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   821
  shows "x \<in> fset S \<Longrightarrow> 0 < fcard S"
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   822
  by (descending) (auto simp add: fcard_raw_def card_gt_0_iff)
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   823
  
36465
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   824
lemma fcard_not_fin:
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   825
  shows "(x |\<notin>| S) = (fcard (finsert x S) = Suc (fcard S))"
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   826
  by (descending) (auto simp add: memb_def fcard_raw_def insert_absorb)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   827
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   828
lemma fcard_suc: "fcard S = Suc n \<Longrightarrow> \<exists>x T. x |\<notin>| T \<and> S = finsert x T \<and> fcard T = n"
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   829
  apply descending
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   830
  apply(simp add: fcard_raw_def memb_def)
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   831
  apply(drule card_eq_SucD)
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   832
  apply(auto)
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   833
  apply(rule_tac x="b" in exI)
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   834
  apply(rule_tac x="removeAll b S" in exI)
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   835
  apply(auto)
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   836
  done
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   837
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   838
lemma fcard_delete:
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   839
  "fcard (fdelete y S) = (if y |\<in>| S then fcard S - 1 else fcard S)"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   840
  by (lifting fcard_raw_delete)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   841
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   842
lemma fcard_suc_memb: 
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   843
  shows "fcard A = Suc n \<Longrightarrow> \<exists>a. a |\<in>| A"
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   844
  apply(descending)
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   845
  apply(simp add: fcard_raw_def memb_def)
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   846
  apply(drule card_eq_SucD)
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   847
  apply(auto)
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   848
  done
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   849
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   850
lemma fin_fcard_not_0: 
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   851
  shows "a |\<in>| A \<Longrightarrow> fcard A \<noteq> 0"
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   852
  by (descending) (auto simp add: fcard_raw_def memb_def)
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   853
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   854
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   855
text {* funion *}
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   856
36352
f71978e47cd5 add bounded_lattice_bot and bounded_lattice_top type classes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36280
diff changeset
   857
lemmas [simp] =
36465
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   858
  sup_bot_left[where 'a="'a fset", standard]
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   859
  sup_bot_right[where 'a="'a fset", standard]
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   860
36352
f71978e47cd5 add bounded_lattice_bot and bounded_lattice_top type classes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36280
diff changeset
   861
lemma funion_finsert[simp]:
f71978e47cd5 add bounded_lattice_bot and bounded_lattice_top type classes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36280
diff changeset
   862
  shows "finsert x S |\<union>| T = finsert x (S |\<union>| T)"
f71978e47cd5 add bounded_lattice_bot and bounded_lattice_top type classes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36280
diff changeset
   863
  by (lifting append.simps(2))
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   864
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   865
lemma singleton_union_left:
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   866
  shows "{|a|} |\<union>| S = finsert a S"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   867
  by simp
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   868
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   869
lemma singleton_union_right:
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   870
  shows "S |\<union>| {|a|} = finsert a S"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   871
  by (subst sup.commute) simp
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   872
39994
7bd8013b903f FSet: give names to respectfulness theorems, rename list_all2_refl to avoid clash
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39302
diff changeset
   873
7bd8013b903f FSet: give names to respectfulness theorems, rename list_all2_refl to avoid clash
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39302
diff changeset
   874
section {* Induction and Cases rules for fsets *}
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   875
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   876
lemma fset_strong_cases:
36465
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   877
  obtains "xs = {||}"
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   878
    | x ys where "x |\<notin>| ys" and "xs = finsert x ys"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   879
  by (lifting fset_raw_strong_cases)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   880
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   881
lemma fset_exhaust[case_names fempty finsert, cases type: fset]:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   882
  shows "\<lbrakk>S = {||} \<Longrightarrow> P; \<And>x S'. S = finsert x S' \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   883
  by (lifting list.exhaust)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   884
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   885
lemma fset_induct_weak[case_names fempty finsert]:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   886
  shows "\<lbrakk>P {||}; \<And>x S. P S \<Longrightarrow> P (finsert x S)\<rbrakk> \<Longrightarrow> P S"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   887
  by (lifting list.induct)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   888
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   889
lemma fset_induct[case_names fempty finsert, induct type: fset]:
36465
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
   890
  assumes prem1: "P {||}"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   891
  and     prem2: "\<And>x S. \<lbrakk>x |\<notin>| S; P S\<rbrakk> \<Longrightarrow> P (finsert x S)"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   892
  shows "P S"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   893
proof(induct S rule: fset_induct_weak)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   894
  case fempty
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   895
  show "P {||}" by (rule prem1)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   896
next
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   897
  case (finsert x S)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   898
  have asm: "P S" by fact
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   899
  show "P (finsert x S)"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   900
    by (cases "x |\<in>| S") (simp_all add: asm prem2)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   901
qed
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   902
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   903
lemma fset_induct2:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   904
  "P {||} {||} \<Longrightarrow>
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   905
  (\<And>x xs. x |\<notin>| xs \<Longrightarrow> P (finsert x xs) {||}) \<Longrightarrow>
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   906
  (\<And>y ys. y |\<notin>| ys \<Longrightarrow> P {||} (finsert y ys)) \<Longrightarrow>
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   907
  (\<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>
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   908
  P xsa ysa"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   909
  apply (induct xsa arbitrary: ysa)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   910
  apply (induct_tac x rule: fset_induct)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   911
  apply simp_all
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   912
  apply (induct_tac xa rule: fset_induct)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   913
  apply simp_all
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   914
  done
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   915
36675
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
   916
lemma fset_fcard_induct:
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
   917
  assumes a: "P {||}"
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
   918
  and     b: "\<And>xs ys. Suc (fcard xs) = (fcard ys) \<Longrightarrow> P xs \<Longrightarrow> P ys"
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
   919
  shows "P zs"
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
   920
proof (induct zs)
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
   921
  show "P {||}" by (rule a)
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
   922
next
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
   923
  fix x :: 'a and zs :: "'a fset"
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
   924
  assume h: "P zs"
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
   925
  assume "x |\<notin>| zs"
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
   926
  then have H1: "Suc (fcard zs) = fcard (finsert x zs)" using fcard_suc by auto
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
   927
  then show "P (finsert x zs)" using b h by simp
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
   928
qed
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
   929
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   930
text {* fmap *}
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   931
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   932
lemma fmap_simps[simp]:
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   933
  fixes f::"'a \<Rightarrow> 'b"
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   934
  shows "fmap f {||} = {||}"
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   935
  and   "fmap f (finsert x S) = finsert (f x) (fmap f S)"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   936
  by (lifting map.simps)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   937
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   938
lemma fmap_set_image:
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   939
  "fset (fmap f S) = f ` (fset S)"
36646
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
   940
  by (induct S) simp_all
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   941
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   942
lemma inj_fmap_eq_iff:
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   943
  "inj f \<Longrightarrow> fmap f S = fmap f T \<longleftrightarrow> S = T"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   944
  by (lifting inj_map_eq_iff)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   945
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   946
lemma fmap_funion: 
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   947
  shows "fmap f (S |\<union>| T) = fmap f S |\<union>| fmap f T"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   948
  by (lifting map_append)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   949
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   950
lemma fin_funion:
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   951
  shows "x |\<in>| S |\<union>| T \<longleftrightarrow> x |\<in>| S \<or> x |\<in>| T"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   952
  by (lifting memb_append)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   953
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   954
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   955
section {* fset *}
36646
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
   956
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   957
lemma fin_set: 
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   958
  shows "x |\<in>| xs \<longleftrightarrow> x \<in> fset xs"
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   959
  by (lifting memb_def)
36646
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
   960
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   961
lemma fnotin_set: 
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   962
  shows "x |\<notin>| xs \<longleftrightarrow> x \<notin> fset xs"
36646
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
   963
  by (simp add: fin_set)
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
   964
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   965
lemma fcard_set: 
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   966
  shows "fcard xs = card (fset xs)"
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   967
  by (lifting fcard_raw_def)
36646
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
   968
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   969
lemma fsubseteq_set: 
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   970
  shows "xs |\<subseteq>| ys \<longleftrightarrow> fset xs \<subseteq> fset ys"
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   971
  by (lifting sub_list_def)
36646
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
   972
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
   973
lemma fsubset_set: 
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   974
  shows "xs |\<subset>| ys \<longleftrightarrow> fset xs \<subset> fset ys"
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   975
  unfolding less_fset_def 
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   976
  by (descending) (auto simp add: sub_list_def)
36646
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
   977
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   978
lemma ffilter_set [simp]: 
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   979
  shows "fset (ffilter P xs) = P \<inter> fset xs"
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   980
  by (descending) (auto simp add: mem_def)
36646
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
   981
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   982
lemma fdelete_set [simp]: 
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   983
  shows "fset (fdelete x xs) = fset xs - {x}"
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   984
  by (lifting set_removeAll)
36646
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
   985
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   986
lemma finter_set [simp]: 
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   987
  shows "fset (xs |\<inter>| ys) = fset xs \<inter> fset ys"
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   988
  by (lifting set_finter_raw)
36646
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
   989
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   990
lemma funion_set [simp]: 
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   991
  shows "fset (xs |\<union>| ys) = fset xs \<union> fset ys"
36646
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
   992
  by (lifting set_append)
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
   993
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   994
lemma fminus_set [simp]: 
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   995
  shows "fset (xs - ys) = fset xs - fset ys"
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
   996
  by (lifting set_fminus_raw)
36675
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
   997
36646
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
   998
lemmas fset_to_set_trans =
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
   999
  fin_set fnotin_set fcard_set fsubseteq_set fsubset_set
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1000
  finter_set funion_set ffilter_set fset_simps
36675
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
  1001
  fset_cong fdelete_set fmap_set_image fminus_set
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
  1002
36646
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
  1003
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1004
text {* ffold *}
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1005
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1006
lemma ffold_nil: 
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1007
  shows "ffold f z {||} = z"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1008
  by (lifting ffold_raw.simps(1)[where 'a="'b" and 'b="'a"])
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1009
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1010
lemma ffold_finsert: "ffold f z (finsert a A) =
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1011
  (if rsp_fold f then if a |\<in>| A then ffold f z A else f a (ffold f z A) else z)"
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1012
  by (descending) (simp add: memb_def)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1013
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1014
lemma fin_commute_ffold:
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1015
  "\<lbrakk>rsp_fold f; h |\<in>| b\<rbrakk> \<Longrightarrow> ffold f z b = f h (ffold f z (fdelete h b))"
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1016
  by (descending) (simp add: memb_def memb_commute_ffold_raw)
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1017
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1018
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1019
text {* fdelete *}
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1020
36465
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
  1021
lemma fin_fdelete:
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1022
  shows "x |\<in>| fdelete y S \<longleftrightarrow> x |\<in>| S \<and> x \<noteq> y"
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1023
  by (descending) (simp add: memb_def)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1024
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1025
lemma fnotin_fdelete:
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1026
  shows "x |\<notin>| fdelete x S"
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1027
  by (descending) (simp add: memb_def)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1028
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1029
lemma fnotin_fdelete_ident:
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1030
  shows "x |\<notin>| S \<Longrightarrow> fdelete x S = S"
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1031
  by (descending) (simp add: memb_def)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1032
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1033
lemma fset_fdelete_cases:
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1034
  shows "S = {||} \<or> (\<exists>x. x |\<in>| S \<and> S = finsert x (fdelete x S))"
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1035
  by (lifting fset_raw_removeAll_cases)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1036
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1037
text {* finite intersection *}
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1038
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1039
lemma finter_empty_l:
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1040
  shows "{||} |\<inter>| S = {||}"
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1041
  by simp
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1042
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1043
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1044
lemma finter_empty_r:
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1045
  shows "S |\<inter>| {||} = {||}"
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1046
  by simp
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1047
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1048
lemma finter_finsert:
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1049
  shows "finsert x S |\<inter>| T = (if x |\<in>| T then finsert x (S |\<inter>| T) else S |\<inter>| T)"
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1050
  by (descending) (simp add: memb_def)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1051
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1052
lemma fin_finter:
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1053
  shows "x |\<in>| (S |\<inter>| T) \<longleftrightarrow> x |\<in>| S \<and> x |\<in>| T"
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1054
  by (descending) (simp add: memb_def)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1055
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1056
lemma fsubset_finsert:
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1057
  shows "finsert x xs |\<subseteq>| ys \<longleftrightarrow> x |\<in>| ys \<and> xs |\<subseteq>| ys"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1058
  by (lifting sub_list_cons)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1059
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1060
lemma 
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1061
  shows "xs |\<subseteq>| ys \<equiv> \<forall>x. x |\<in>| xs \<longrightarrow> x |\<in>| ys"
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1062
  by (descending) (auto simp add: sub_list_def memb_def)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1063
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1064
lemma fsubset_fin: 
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1065
  shows "xs |\<subseteq>| ys = (\<forall>x. x |\<in>| xs \<longrightarrow> x |\<in>| ys)"
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1066
  by (descending) (auto simp add: sub_list_def memb_def)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1067
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1068
lemma fminus_fin: 
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1069
  shows "x |\<in>| xs - ys \<longleftrightarrow> x |\<in>| xs \<and> x |\<notin>| ys"
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1070
  by (descending) (simp add: memb_def)
36675
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
  1071
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1072
lemma fminus_red: 
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1073
  shows "finsert x xs - ys = (if x |\<in>| ys then xs - ys else finsert x (xs - ys))"
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1074
  by (descending) (auto simp add: memb_def)
36675
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
  1075
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1076
lemma fminus_red_fin [simp]: 
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1077
  shows "x |\<in>| ys \<Longrightarrow> finsert x xs - ys = xs - ys"
36675
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
  1078
  by (simp add: fminus_red)
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
  1079
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1080
lemma fminus_red_fnotin[simp]: 
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1081
  shows "x |\<notin>| ys \<Longrightarrow> finsert x xs - ys = finsert x (xs - ys)"
36675
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
  1082
  by (simp add: fminus_red)
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
  1083
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1084
lemma fset_eq_iff:
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1085
  shows "S = T \<longleftrightarrow> (\<forall>x. (x |\<in>| S) = (x |\<in>| T))"
39995
849578dd6127 FSet tuned
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39994
diff changeset
  1086
  by (descending) (auto simp add: memb_def)
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1087
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1088
(* We cannot write it as "assumes .. shows" since Isabelle changes
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1089
   the quantifiers to schematic variables and reintroduces them in
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1090
   a different order *)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1091
lemma fset_eq_cases:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1092
 "\<lbrakk>a1 = a2;
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1093
   \<And>a b xs. \<lbrakk>a1 = finsert a (finsert b xs); a2 = finsert b (finsert a xs)\<rbrakk> \<Longrightarrow> P;
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1094
   \<lbrakk>a1 = {||}; a2 = {||}\<rbrakk> \<Longrightarrow> P; \<And>xs ys. \<lbrakk>a1 = ys; a2 = xs; xs = ys\<rbrakk> \<Longrightarrow> P;
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1095
   \<And>a xs. \<lbrakk>a1 = finsert a (finsert a xs); a2 = finsert a xs\<rbrakk> \<Longrightarrow> P;
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1096
   \<And>xs ys a. \<lbrakk>a1 = finsert a xs; a2 = finsert a ys; xs = ys\<rbrakk> \<Longrightarrow> P;
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1097
   \<And>xs1 xs2 xs3. \<lbrakk>a1 = xs1; a2 = xs3; xs1 = xs2; xs2 = xs3\<rbrakk> \<Longrightarrow> P\<rbrakk>
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1098
  \<Longrightarrow> P"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1099
  by (lifting list_eq2.cases[simplified list_eq2_equiv[symmetric]])
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1100
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1101
lemma fset_eq_induct:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1102
  assumes "x1 = x2"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1103
  and "\<And>a b xs. P (finsert a (finsert b xs)) (finsert b (finsert a xs))"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1104
  and "P {||} {||}"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1105
  and "\<And>xs ys. \<lbrakk>xs = ys; P xs ys\<rbrakk> \<Longrightarrow> P ys xs"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1106
  and "\<And>a xs. P (finsert a (finsert a xs)) (finsert a xs)"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1107
  and "\<And>xs ys a. \<lbrakk>xs = ys; P xs ys\<rbrakk> \<Longrightarrow> P (finsert a xs) (finsert a ys)"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1108
  and "\<And>xs1 xs2 xs3. \<lbrakk>xs1 = xs2; P xs1 xs2; xs2 = xs3; P xs2 xs3\<rbrakk> \<Longrightarrow> P xs1 xs3"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1109
  shows "P x1 x2"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1110
  using assms
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1111
  by (lifting list_eq2.induct[simplified list_eq2_equiv[symmetric]])
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1112
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1113
section {* fconcat *}
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1114
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1115
lemma fconcat_empty:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1116
  shows "fconcat {||} = {||}"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1117
  by (lifting concat.simps(1))
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1118
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1119
lemma fconcat_insert:
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1120
  shows "fconcat (finsert x S) = x |\<union>| fconcat S"
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1121
  by (lifting concat.simps(2))
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1122
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1123
lemma 
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1124
  shows "fconcat (xs |\<union>| ys) = fconcat xs |\<union>| fconcat ys"
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1125
  by (lifting concat_append)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1126
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1127
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1128
section {* ffilter *}
36639
6243b49498ea added function ffilter and some lemmas from Finite_Set to the FSet theory
bulwahn
parents: 36524
diff changeset
  1129
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1130
lemma subseteq_filter: 
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1131
  shows "ffilter P xs <= ffilter Q xs = (\<forall> x. x |\<in>| xs \<longrightarrow> P x \<longrightarrow> Q x)"
39994
7bd8013b903f FSet: give names to respectfulness theorems, rename list_all2_refl to avoid clash
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39302
diff changeset
  1132
  by  (descending) (auto simp add: memb_def sub_list_def)
36639
6243b49498ea added function ffilter and some lemmas from Finite_Set to the FSet theory
bulwahn
parents: 36524
diff changeset
  1133
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1134
lemma eq_ffilter: 
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1135
  shows "(ffilter P xs = ffilter Q xs) = (\<forall>x. x |\<in>| xs \<longrightarrow> P x = Q x)"
39994
7bd8013b903f FSet: give names to respectfulness theorems, rename list_all2_refl to avoid clash
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39302
diff changeset
  1136
  by (descending) (auto simp add: memb_def)
36639
6243b49498ea added function ffilter and some lemmas from Finite_Set to the FSet theory
bulwahn
parents: 36524
diff changeset
  1137
39995
849578dd6127 FSet tuned
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39994
diff changeset
  1138
lemma subset_ffilter:
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1139
  shows "(\<And>x. x |\<in>| xs \<Longrightarrow> P x \<Longrightarrow> Q x) \<Longrightarrow> (x |\<in>| xs & \<not> P x & Q x) \<Longrightarrow> ffilter P xs < ffilter Q xs"
39995
849578dd6127 FSet tuned
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39994
diff changeset
  1140
  unfolding less_fset_def by (auto simp add: subseteq_filter eq_ffilter)
849578dd6127 FSet tuned
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39994
diff changeset
  1141
36639
6243b49498ea added function ffilter and some lemmas from Finite_Set to the FSet theory
bulwahn
parents: 36524
diff changeset
  1142
6243b49498ea added function ffilter and some lemmas from Finite_Set to the FSet theory
bulwahn
parents: 36524
diff changeset
  1143
section {* lemmas transferred from Finite_Set theory *}
6243b49498ea added function ffilter and some lemmas from Finite_Set to the FSet theory
bulwahn
parents: 36524
diff changeset
  1144
6243b49498ea added function ffilter and some lemmas from Finite_Set to the FSet theory
bulwahn
parents: 36524
diff changeset
  1145
text {* finiteness for finite sets holds *}
39994
7bd8013b903f FSet: give names to respectfulness theorems, rename list_all2_refl to avoid clash
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39302
diff changeset
  1146
lemma finite_fset [simp]: 
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1147
  shows "finite (fset S)"
36646
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
  1148
  by (induct S) auto
36639
6243b49498ea added function ffilter and some lemmas from Finite_Set to the FSet theory
bulwahn
parents: 36524
diff changeset
  1149
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1150
lemma fset_choice: 
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1151
  shows "\<forall>x. x |\<in>| A \<longrightarrow> (\<exists>y. P x y) \<Longrightarrow> \<exists>f. \<forall>x. x |\<in>| A \<longrightarrow> P x (f x)"
36646
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
  1152
  unfolding fset_to_set_trans
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
  1153
  by (rule finite_set_choice[simplified Ball_def, OF finite_fset])
36639
6243b49498ea added function ffilter and some lemmas from Finite_Set to the FSet theory
bulwahn
parents: 36524
diff changeset
  1154
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1155
lemma fsubseteq_fempty:
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1156
  shows "xs |\<subseteq>| {||} \<longleftrightarrow> xs = {||}"
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1157
  by (metis finter_empty_r le_iff_inf)
36646
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
  1158
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1159
lemma not_fsubset_fnil: 
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1160
  shows "\<not> xs |\<subset>| {||}"
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1161
  by (metis fset_simps(1) fsubset_set not_psubset_empty)
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1162
  
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1163
lemma fcard_mono: 
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1164
  shows "xs |\<subseteq>| ys \<Longrightarrow> fcard xs \<le> fcard ys"
36646
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
  1165
  unfolding fset_to_set_trans
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
  1166
  by (rule card_mono[OF finite_fset])
36639
6243b49498ea added function ffilter and some lemmas from Finite_Set to the FSet theory
bulwahn
parents: 36524
diff changeset
  1167
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1168
lemma fcard_fseteq: 
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1169
  shows "xs |\<subseteq>| ys \<Longrightarrow> fcard ys \<le> fcard xs \<Longrightarrow> xs = ys"
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1170
  unfolding fcard_set fsubseteq_set
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1171
  by (simp add: card_seteq[OF finite_fset] fset_cong)
36646
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
  1172
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1173
lemma psubset_fcard_mono: 
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1174
  shows "xs |\<subset>| ys \<Longrightarrow> fcard xs < fcard ys"
36646
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
  1175
  unfolding fset_to_set_trans
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
  1176
  by (rule psubset_card_mono[OF finite_fset])
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
  1177
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1178
lemma fcard_funion_finter: 
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1179
  shows "fcard xs + fcard ys = fcard (xs |\<union>| ys) + fcard (xs |\<inter>| ys)"
36646
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
  1180
  unfolding fset_to_set_trans
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
  1181
  by (rule card_Un_Int[OF finite_fset finite_fset])
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
  1182
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1183
lemma fcard_funion_disjoint: 
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1184
  shows "xs |\<inter>| ys = {||} \<Longrightarrow> fcard (xs |\<union>| ys) = fcard xs + fcard ys"
36646
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
  1185
  unfolding fset_to_set_trans
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
  1186
  by (rule card_Un_disjoint[OF finite_fset finite_fset])
36639
6243b49498ea added function ffilter and some lemmas from Finite_Set to the FSet theory
bulwahn
parents: 36524
diff changeset
  1187
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1188
lemma fcard_delete1_less: 
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1189
  shows "x |\<in>| xs \<Longrightarrow> fcard (fdelete x xs) < fcard xs"
36646
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
  1190
  unfolding fset_to_set_trans
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
  1191
  by (rule card_Diff1_less[OF finite_fset])
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
  1192
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1193
lemma fcard_delete2_less: 
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1194
  shows "x |\<in>| xs \<Longrightarrow> y |\<in>| xs \<Longrightarrow> fcard (fdelete y (fdelete x xs)) < fcard xs"
36646
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
  1195
  unfolding fset_to_set_trans
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
  1196
  by (rule card_Diff2_less[OF finite_fset])
36639
6243b49498ea added function ffilter and some lemmas from Finite_Set to the FSet theory
bulwahn
parents: 36524
diff changeset
  1197
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1198
lemma fcard_delete1_le: 
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1199
  shows "fcard (fdelete x xs) \<le> fcard xs"
36646
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
  1200
  unfolding fset_to_set_trans
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
  1201
  by (rule card_Diff1_le[OF finite_fset])
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
  1202
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1203
lemma fcard_psubset: 
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1204
  shows "ys |\<subseteq>| xs \<Longrightarrow> fcard ys < fcard xs \<Longrightarrow> ys |\<subset>| xs"
36646
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
  1205
  unfolding fset_to_set_trans
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
  1206
  by (rule card_psubset[OF finite_fset])
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
  1207
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1208
lemma fcard_fmap_le: 
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1209
  shows "fcard (fmap f xs) \<le> fcard xs"
36646
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
  1210
  unfolding fset_to_set_trans
8687bc6608d6 Translating lemmas from Finite_Set to FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36639
diff changeset
  1211
  by (rule card_image_le[OF finite_fset])
36639
6243b49498ea added function ffilter and some lemmas from Finite_Set to the FSet theory
bulwahn
parents: 36524
diff changeset
  1212
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1213
lemma fin_fminus_fnotin: 
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1214
  shows "x |\<in>| F - S \<Longrightarrow> x |\<notin>| S"
36675
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
  1215
  unfolding fset_to_set_trans
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
  1216
  by blast
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
  1217
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1218
lemma fin_fnotin_fminus: 
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1219
  shows "x |\<in>| S \<Longrightarrow> x |\<notin>| F - S"
36675
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
  1220
  unfolding fset_to_set_trans
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
  1221
  by blast
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
  1222
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1223
lemma fin_mdef: 
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1224
  "x |\<in>| F \<longleftrightarrow> x |\<notin>| (F - {|x|}) \<and> F = finsert x (F - {|x|})"
36675
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
  1225
  unfolding fset_to_set_trans
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
  1226
  by blast
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
  1227
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
  1228
lemma fcard_fminus_finsert[simp]:
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
  1229
  assumes "a |\<in>| A" and "a |\<notin>| B"
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
  1230
  shows "fcard(A - finsert a B) = fcard(A - B) - 1"
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1231
  using assms 
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1232
  unfolding fset_to_set_trans
36675
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
  1233
  by (rule card_Diff_insert[OF finite_fset])
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
  1234
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
  1235
lemma fcard_fminus_fsubset:
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
  1236
  assumes "B |\<subseteq>| A"
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
  1237
  shows "fcard (A - B) = fcard A - fcard B"
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
  1238
  using assms unfolding fset_to_set_trans
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
  1239
  by (rule card_Diff_subset[OF finite_fset])
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
  1240
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
  1241
lemma fcard_fminus_subset_finter:
39996
c02078ff8691 FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 39995
diff changeset
  1242
  shows "fcard (A - B) = fcard A - fcard (A |\<inter>| B)"
36675
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
  1243
  unfolding fset_to_set_trans
37634
2116425cebc8 cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents: 37492
diff changeset
  1244
  by (rule card_Diff_subset_Int) (fold finter_set, rule finite_fset)
36675
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
  1245
806ea6e282e4 fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36646
diff changeset
  1246
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1247
ML {*
36465
15e834a03509 Tuned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36352
diff changeset
  1248
fun dest_fsetT (Type (@{type_name fset}, [T])) = T
36280
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1249
  | dest_fsetT T = raise TYPE ("dest_fsetT: fset type expected", [T], []);
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1250
*}
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1251
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1252
no_notation
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1253
  list_eq (infix "\<approx>" 50)
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1254
c4f5823f282d Finite set theory
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
  1255
end