src/HOL/Finite.ML
author paulson
Fri, 30 May 1997 15:17:14 +0200
changeset 3368 be517d000c02
parent 3352 04502e5431fb
child 3382 8db8fc8607d9
permissions -rw-r--r--
Many new theorems about cardinality
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
     1
(*  Title:      HOL/Finite.thy
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
     3
    Author:     Lawrence C Paulson & Tobias Nipkow
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
     4
    Copyright   1995  University of Cambridge & TU Muenchen
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     5
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
     6
Finite sets and their cardinality
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     7
*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     8
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     9
open Finite;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    10
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
    11
section "The finite powerset operator -- Fin";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    12
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    13
goalw Finite.thy Fin.defs "!!A B. A<=B ==> Fin(A) <= Fin(B)";
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
    14
by (rtac lfp_mono 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    15
by (REPEAT (ares_tac basic_monos 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    16
qed "Fin_mono";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    17
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    18
goalw Finite.thy Fin.defs "Fin(A) <= Pow(A)";
2922
580647a879cf Using Blast_tac
paulson
parents: 2031
diff changeset
    19
by (blast_tac (!claset addSIs [lfp_lowerbound]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    20
qed "Fin_subset_Pow";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    21
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    22
(* A : Fin(B) ==> A <= B *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    23
val FinD = Fin_subset_Pow RS subsetD RS PowD;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    24
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    25
(*Discharging ~ x:y entails extra work*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    26
val major::prems = goal Finite.thy 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    27
    "[| F:Fin(A);  P({}); \
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
    28
\       !!F x. [| x:A;  F:Fin(A);  x~:F;  P(F) |] ==> P(insert x F) \
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    29
\    |] ==> P(F)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    30
by (rtac (major RS Fin.induct) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    31
by (excluded_middle_tac "a:b" 2);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    32
by (etac (insert_absorb RS ssubst) 3 THEN assume_tac 3);   (*backtracking!*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    33
by (REPEAT (ares_tac prems 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    34
qed "Fin_induct";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    35
1264
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 923
diff changeset
    36
Addsimps Fin.intrs;
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    37
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    38
(*The union of two finite sets is finite*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    39
val major::prems = goal Finite.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    40
    "[| F: Fin(A);  G: Fin(A) |] ==> F Un G : Fin(A)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    41
by (rtac (major RS Fin_induct) 1);
1264
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 923
diff changeset
    42
by (ALLGOALS (asm_simp_tac (!simpset addsimps (prems @ [Un_insert_left]))));
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    43
qed "Fin_UnI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    44
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    45
(*Every subset of a finite set is finite*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    46
val [subs,fin] = goal Finite.thy "[| A<=B;  B: Fin(M) |] ==> A: Fin(M)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    47
by (EVERY1 [subgoal_tac "ALL C. C<=B --> C: Fin(M)",
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
    48
            rtac mp, etac spec,
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
    49
            rtac subs]);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    50
by (rtac (fin RS Fin_induct) 1);
1264
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 923
diff changeset
    51
by (simp_tac (!simpset addsimps [subset_Un_eq]) 1);
1786
8a31d85d27b8 best_tac, deepen_tac and safe_tac now also use default claset.
berghofe
parents: 1782
diff changeset
    52
by (safe_tac (!claset addSDs [subset_insert_iff RS iffD1]));
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    53
by (eres_inst_tac [("t","C")] (insert_Diff RS subst) 2);
1264
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 923
diff changeset
    54
by (ALLGOALS Asm_simp_tac);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    55
qed "Fin_subset";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    56
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    57
goal Finite.thy "(F Un G : Fin(A)) = (F: Fin(A) & G: Fin(A))";
2922
580647a879cf Using Blast_tac
paulson
parents: 2031
diff changeset
    58
by (blast_tac (!claset addIs [Fin_UnI] addDs
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    59
                [Un_upper1 RS Fin_subset, Un_upper2 RS Fin_subset]) 1);
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    60
qed "subset_Fin";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    61
Addsimps[subset_Fin];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    62
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    63
goal Finite.thy "(insert a A : Fin M) = (a:M & A : Fin M)";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
    64
by (stac insert_is_Un 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
    65
by (Simp_tac 1);
2922
580647a879cf Using Blast_tac
paulson
parents: 2031
diff changeset
    66
by (blast_tac (!claset addSIs Fin.intrs addDs [FinD]) 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    67
qed "insert_Fin";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    68
Addsimps[insert_Fin];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    69
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    70
(*The image of a finite set is finite*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    71
val major::_ = goal Finite.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    72
    "F: Fin(A) ==> h``F : Fin(h``A)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    73
by (rtac (major RS Fin_induct) 1);
1264
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 923
diff changeset
    74
by (Simp_tac 1);
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 923
diff changeset
    75
by (asm_simp_tac
1660
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
    76
    (!simpset addsimps [image_eqI RS Fin.insertI, image_insert]
2031
03a843f0f447 Ran expandshort
paulson
parents: 1786
diff changeset
    77
              delsimps [insert_Fin]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    78
qed "Fin_imageI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    79
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    80
val major::prems = goal Finite.thy 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
    81
    "[| c: Fin(A);  b: Fin(A);                                  \
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
    82
\       P(b);                                                   \
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    83
\       !!(x::'a) y. [| x:A; y: Fin(A);  x:y;  P(y) |] ==> P(y-{x}) \
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    84
\    |] ==> c<=b --> P(b-c)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    85
by (rtac (major RS Fin_induct) 1);
2031
03a843f0f447 Ran expandshort
paulson
parents: 1786
diff changeset
    86
by (stac Diff_insert 2);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    87
by (ALLGOALS (asm_simp_tac
1264
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 923
diff changeset
    88
                (!simpset addsimps (prems@[Diff_subset RS Fin_subset]))));
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    89
val lemma = result();
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    90
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    91
val prems = goal Finite.thy 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
    92
    "[| b: Fin(A);                                              \
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
    93
\       P(b);                                                   \
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    94
\       !!x y. [| x:A; y: Fin(A);  x:y;  P(y) |] ==> P(y-{x}) \
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    95
\    |] ==> P({})";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    96
by (rtac (Diff_cancel RS subst) 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    97
by (rtac (lemma RS mp) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    98
by (REPEAT (ares_tac (subset_refl::prems) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    99
qed "Fin_empty_induct";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   100
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   101
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   102
section "The predicate 'finite'";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   103
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   104
val major::prems = goalw Finite.thy [finite_def]
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   105
    "[| finite F;  P({}); \
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   106
\       !!F x. [| finite F;  x~:F;  P(F) |] ==> P(insert x F) \
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   107
\    |] ==> P(F)";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   108
by (rtac (major RS Fin_induct) 1);
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   109
by (REPEAT (ares_tac prems 1));
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   110
qed "finite_induct";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   111
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   112
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   113
goalw Finite.thy [finite_def] "finite {}";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   114
by (Simp_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   115
qed "finite_emptyI";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   116
Addsimps [finite_emptyI];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   117
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   118
goalw Finite.thy [finite_def] "!!A. finite A ==> finite(insert a A)";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   119
by (Asm_simp_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   120
qed "finite_insertI";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   121
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   122
(*The union of two finite sets is finite*)
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   123
goalw Finite.thy [finite_def]
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   124
    "!!F. [| finite F;  finite G |] ==> finite(F Un G)";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   125
by (Asm_simp_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   126
qed "finite_UnI";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   127
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   128
goalw Finite.thy [finite_def] "!!A. [| A<=B;  finite B |] ==> finite A";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   129
by (etac Fin_subset 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   130
by (assume_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   131
qed "finite_subset";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   132
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   133
goalw Finite.thy [finite_def] "finite(F Un G) = (finite F & finite G)";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   134
by (Simp_tac 1);
3352
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   135
qed "finite_Un_eq";
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   136
Addsimps[finite_Un_eq];
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   137
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   138
goalw Finite.thy [finite_def] "finite(insert a A) = finite(A)";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   139
by (Simp_tac 1);
3352
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   140
qed "finite_insert";
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   141
Addsimps[finite_insert];
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   142
1618
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1553
diff changeset
   143
(* finite B ==> finite (B - Ba) *)
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1553
diff changeset
   144
bind_thm ("finite_Diff", Diff_subset RS finite_subset);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   145
Addsimps [finite_Diff];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   146
3368
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   147
goal Finite.thy "finite(A-{a}) = finite(A)";
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   148
by (case_tac "a:A" 1);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   149
br (finite_insert RS sym RS trans) 1;
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   150
by (stac insert_Diff 1);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   151
by (ALLGOALS Asm_simp_tac);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   152
qed "finite_Diff_singleton";
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   153
AddIffs [finite_Diff_singleton];
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   154
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   155
(*The image of a finite set is finite*)
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   156
goal Finite.thy "!!F. finite F ==> finite(h``F)";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   157
by (etac finite_induct 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   158
by (ALLGOALS Asm_simp_tac);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   159
qed "finite_imageI";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   160
3368
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   161
goal Finite.thy "!!A. finite B ==> !A. B = f``A --> inj_onto f A --> finite A";
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   162
by (etac finite_induct 1);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   163
by (ALLGOALS Asm_simp_tac);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   164
by (Step_tac 1);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   165
by (subgoal_tac "A={}" 1);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   166
by (blast_tac (!claset addSEs [equalityE]) 2);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   167
by (subgoal_tac "EX y:A. f y = x & F = f``(A-{y})" 2);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   168
by (Step_tac 1);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   169
bw inj_onto_def;
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   170
by (Blast_tac 2);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   171
by (ALLGOALS Asm_simp_tac);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   172
by (thin_tac "ALL A. ?PP(A)" 1);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   173
by (forward_tac [[equalityD1, insertI1] MRS subsetD] 1);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   174
by (Step_tac 1);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   175
by (res_inst_tac [("x","xa")] bexI 1);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   176
by (ALLGOALS Asm_simp_tac);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   177
be equalityE 1;
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   178
br equalityI 1;
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   179
by (Blast_tac 2);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   180
by (Best_tac 1);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   181
val lemma = result();
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   182
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   183
goal Finite.thy "!!A. [| finite(f``A);  inj_onto f A |] ==> finite A";
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   184
bd lemma 1;
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   185
by (Blast_tac 1);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   186
qed "finite_imageD";
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   187
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   188
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   189
(** The powerset of a finite set **)
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   190
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   191
goal Finite.thy "!!A. finite(Pow A) ==> finite A";
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   192
by (subgoal_tac "finite ((%x.{x})``A)" 1);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   193
br finite_subset 2;
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   194
ba 3;
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   195
by (ALLGOALS
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   196
    (fast_tac (!claset addSDs [rewrite_rule [inj_onto_def] finite_imageD])));
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   197
val lemma = result();
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   198
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   199
goal Finite.thy "finite(Pow A) = finite A";
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   200
br iffI 1;
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   201
be lemma 1;
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   202
(*Opposite inclusion: finite A ==> finite (Pow A) *)
3340
a886795c9dce Two results suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   203
by (etac finite_induct 1);
a886795c9dce Two results suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   204
by (ALLGOALS 
a886795c9dce Two results suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   205
    (asm_simp_tac
a886795c9dce Two results suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   206
     (!simpset addsimps [finite_UnI, finite_imageI, Pow_insert])));
3368
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   207
qed "finite_Pow_iff";
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   208
AddIffs [finite_Pow_iff];
3340
a886795c9dce Two results suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   209
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   210
val major::prems = goalw Finite.thy [finite_def]
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   211
    "[| finite A;                                       \
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   212
\       P(A);                                           \
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   213
\       !!a A. [| finite A;  a:A;  P(A) |] ==> P(A-{a}) \
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   214
\    |] ==> P({})";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   215
by (rtac (major RS Fin_empty_induct) 1);
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   216
by (REPEAT (ares_tac (subset_refl::prems) 1));
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   217
qed "finite_empty_induct";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   218
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   219
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   220
section "Finite cardinality -- 'card'";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   221
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   222
goal Set.thy "{f i |i. P i | i=n} = insert (f n) {f i|i. P i}";
2922
580647a879cf Using Blast_tac
paulson
parents: 2031
diff changeset
   223
by (Blast_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   224
val Collect_conv_insert = result();
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   225
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   226
goalw Finite.thy [card_def] "card {} = 0";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   227
by (rtac Least_equality 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   228
by (ALLGOALS Asm_full_simp_tac);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   229
qed "card_empty";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   230
Addsimps [card_empty];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   231
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   232
val [major] = goal Finite.thy
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   233
  "finite A ==> ? (n::nat) f. A = {f i |i. i<n}";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   234
by (rtac (major RS finite_induct) 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   235
 by (res_inst_tac [("x","0")] exI 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   236
 by (Simp_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   237
by (etac exE 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   238
by (etac exE 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   239
by (hyp_subst_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   240
by (res_inst_tac [("x","Suc n")] exI 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   241
by (res_inst_tac [("x","%i. if i<n then f i else x")] exI 1);
1660
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   242
by (asm_simp_tac (!simpset addsimps [Collect_conv_insert, less_Suc_eq]
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   243
                          addcongs [rev_conj_cong]) 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   244
qed "finite_has_card";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   245
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   246
goal Finite.thy
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   247
  "!!A.[| x ~: A; insert x A = {f i|i.i<n} |] ==> \
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   248
\  ? m::nat. m<n & (? g. A = {g i|i.i<m})";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   249
by (res_inst_tac [("n","n")] natE 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   250
 by (hyp_subst_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   251
 by (Asm_full_simp_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   252
by (rename_tac "m" 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   253
by (hyp_subst_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   254
by (case_tac "? a. a:A" 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   255
 by (res_inst_tac [("x","0")] exI 2);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   256
 by (Simp_tac 2);
2922
580647a879cf Using Blast_tac
paulson
parents: 2031
diff changeset
   257
 by (Blast_tac 2);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   258
by (etac exE 1);
1660
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   259
by (simp_tac (!simpset addsimps [less_Suc_eq]) 1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   260
by (rtac exI 1);
1782
ab45b881fa62 Shortened a proof
paulson
parents: 1760
diff changeset
   261
by (rtac (refl RS disjI2 RS conjI) 1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   262
by (etac equalityE 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   263
by (asm_full_simp_tac
1660
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   264
     (!simpset addsimps [subset_insert,Collect_conv_insert, less_Suc_eq]) 1);
2922
580647a879cf Using Blast_tac
paulson
parents: 2031
diff changeset
   265
by (safe_tac (!claset));
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   266
  by (Asm_full_simp_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   267
  by (res_inst_tac [("x","%i. if f i = f m then a else f i")] exI 1);
1786
8a31d85d27b8 best_tac, deepen_tac and safe_tac now also use default claset.
berghofe
parents: 1782
diff changeset
   268
  by (SELECT_GOAL(safe_tac (!claset))1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   269
   by (subgoal_tac "x ~= f m" 1);
2922
580647a879cf Using Blast_tac
paulson
parents: 2031
diff changeset
   270
    by (Blast_tac 2);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   271
   by (subgoal_tac "? k. f k = x & k<m" 1);
2922
580647a879cf Using Blast_tac
paulson
parents: 2031
diff changeset
   272
    by (Blast_tac 2);
1786
8a31d85d27b8 best_tac, deepen_tac and safe_tac now also use default claset.
berghofe
parents: 1782
diff changeset
   273
   by (SELECT_GOAL(safe_tac (!claset))1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   274
   by (res_inst_tac [("x","k")] exI 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   275
   by (Asm_simp_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   276
  by (simp_tac (!simpset setloop (split_tac [expand_if])) 1);
2922
580647a879cf Using Blast_tac
paulson
parents: 2031
diff changeset
   277
  by (Blast_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   278
 bd sym 1;
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   279
 by (rotate_tac ~1 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   280
 by (Asm_full_simp_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   281
 by (res_inst_tac [("x","%i. if f i = f m then a else f i")] exI 1);
1786
8a31d85d27b8 best_tac, deepen_tac and safe_tac now also use default claset.
berghofe
parents: 1782
diff changeset
   282
 by (SELECT_GOAL(safe_tac (!claset))1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   283
  by (subgoal_tac "x ~= f m" 1);
2922
580647a879cf Using Blast_tac
paulson
parents: 2031
diff changeset
   284
   by (Blast_tac 2);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   285
  by (subgoal_tac "? k. f k = x & k<m" 1);
2922
580647a879cf Using Blast_tac
paulson
parents: 2031
diff changeset
   286
   by (Blast_tac 2);
1786
8a31d85d27b8 best_tac, deepen_tac and safe_tac now also use default claset.
berghofe
parents: 1782
diff changeset
   287
  by (SELECT_GOAL(safe_tac (!claset))1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   288
  by (res_inst_tac [("x","k")] exI 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   289
  by (Asm_simp_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   290
 by (simp_tac (!simpset setloop (split_tac [expand_if])) 1);
2922
580647a879cf Using Blast_tac
paulson
parents: 2031
diff changeset
   291
 by (Blast_tac 1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   292
by (res_inst_tac [("x","%j. if f j = f i then f m else f j")] exI 1);
1786
8a31d85d27b8 best_tac, deepen_tac and safe_tac now also use default claset.
berghofe
parents: 1782
diff changeset
   293
by (SELECT_GOAL(safe_tac (!claset))1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   294
 by (subgoal_tac "x ~= f i" 1);
2922
580647a879cf Using Blast_tac
paulson
parents: 2031
diff changeset
   295
  by (Blast_tac 2);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   296
 by (case_tac "x = f m" 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   297
  by (res_inst_tac [("x","i")] exI 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   298
  by (Asm_simp_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   299
 by (subgoal_tac "? k. f k = x & k<m" 1);
2922
580647a879cf Using Blast_tac
paulson
parents: 2031
diff changeset
   300
  by (Blast_tac 2);
1786
8a31d85d27b8 best_tac, deepen_tac and safe_tac now also use default claset.
berghofe
parents: 1782
diff changeset
   301
 by (SELECT_GOAL(safe_tac (!claset))1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   302
 by (res_inst_tac [("x","k")] exI 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   303
 by (Asm_simp_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   304
by (simp_tac (!simpset setloop (split_tac [expand_if])) 1);
2922
580647a879cf Using Blast_tac
paulson
parents: 2031
diff changeset
   305
by (Blast_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   306
val lemma = result();
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   307
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   308
goal Finite.thy "!!A. [| finite A; x ~: A |] ==> \
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   309
\ (LEAST n. ? f. insert x A = {f i|i.i<n}) = Suc(LEAST n. ? f. A={f i|i.i<n})";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   310
by (rtac Least_equality 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   311
 bd finite_has_card 1;
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   312
 be exE 1;
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   313
 by (dres_inst_tac [("P","%n.? f. A={f i|i.i<n}")] LeastI 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   314
 be exE 1;
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   315
 by (res_inst_tac
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   316
   [("x","%i. if i<(LEAST n. ? f. A={f i |i. i < n}) then f i else x")] exI 1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   317
 by (simp_tac
1660
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   318
    (!simpset addsimps [Collect_conv_insert, less_Suc_eq] 
2031
03a843f0f447 Ran expandshort
paulson
parents: 1786
diff changeset
   319
              addcongs [rev_conj_cong]) 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   320
 be subst 1;
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   321
 br refl 1;
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   322
by (rtac notI 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   323
by (etac exE 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   324
by (dtac lemma 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   325
 ba 1;
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   326
by (etac exE 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   327
by (etac conjE 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   328
by (dres_inst_tac [("P","%x. ? g. A = {g i |i. i < x}")] Least_le 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   329
by (dtac le_less_trans 1 THEN atac 1);
1660
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   330
by (asm_full_simp_tac (!simpset addsimps [less_Suc_eq]) 1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   331
by (etac disjE 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   332
by (etac less_asym 1 THEN atac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   333
by (hyp_subst_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   334
by (Asm_full_simp_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   335
val lemma = result();
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   336
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   337
goalw Finite.thy [card_def]
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   338
  "!!A. [| finite A; x ~: A |] ==> card(insert x A) = Suc(card A)";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   339
by (etac lemma 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   340
by (assume_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   341
qed "card_insert_disjoint";
3352
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   342
Addsimps [card_insert_disjoint];
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   343
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   344
goal Finite.thy  "!!A. finite A ==> !B. B <= A --> card(B) <= card(A)";
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   345
by (etac finite_induct 1);
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   346
by (Simp_tac 1);
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   347
by (strip_tac 1);
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   348
by (case_tac "x:B" 1);
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   349
 by (dtac mk_disjoint_insert 1);
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   350
 by (SELECT_GOAL(safe_tac (!claset))1);
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   351
 by (rotate_tac ~1 1);
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   352
 by (asm_full_simp_tac (!simpset addsimps [subset_insert_iff,finite_subset]) 1);
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   353
by (rotate_tac ~1 1);
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   354
by (asm_full_simp_tac (!simpset addsimps [subset_insert_iff,finite_subset]) 1);
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   355
qed_spec_mp "card_mono";
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   356
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   357
goal Finite.thy "!!A B. [| finite A; finite B |]\
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   358
\                       ==> A Int B = {} --> card(A Un B) = card A + card B";
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   359
by (etac finite_induct 1);
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   360
by (ALLGOALS 
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   361
    (asm_simp_tac (!simpset addsimps [Un_insert_left, Int_insert_left]
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   362
		            setloop split_tac [expand_if])));
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   363
qed_spec_mp "card_Un_disjoint";
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   364
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   365
goal Finite.thy "!!A. [| finite A; B<=A |] ==> card A - card B = card (A - B)";
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   366
by (subgoal_tac "(A-B) Un B = A" 1);
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   367
by (Blast_tac 2);
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   368
br (add_right_cancel RS iffD1) 1;
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   369
br (card_Un_disjoint RS subst) 1;
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   370
be ssubst 4;
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   371
by (Blast_tac 3);
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   372
by (ALLGOALS 
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   373
    (asm_simp_tac
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   374
     (!simpset addsimps [add_commute, not_less_iff_le, 
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   375
			 add_diff_inverse, card_mono, finite_subset])));
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   376
qed "card_Diff_subset";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   377
1618
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1553
diff changeset
   378
goal Finite.thy "!!A. [| finite A; x: A |] ==> Suc(card(A-{x})) = card A";
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1553
diff changeset
   379
by (res_inst_tac [("t", "A")] (insert_Diff RS subst) 1);
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1553
diff changeset
   380
by (assume_tac 1);
3352
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   381
by (Asm_simp_tac 1);
1618
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1553
diff changeset
   382
qed "card_Suc_Diff";
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1553
diff changeset
   383
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1553
diff changeset
   384
goal Finite.thy "!!A. [| finite A; x: A |] ==> card(A-{x}) < card A";
2031
03a843f0f447 Ran expandshort
paulson
parents: 1786
diff changeset
   385
by (rtac Suc_less_SucD 1);
1618
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1553
diff changeset
   386
by (asm_simp_tac (!simpset addsimps [card_Suc_Diff]) 1);
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1553
diff changeset
   387
qed "card_Diff";
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1553
diff changeset
   388
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   389
val [major] = goal Finite.thy
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   390
  "finite A ==> card(insert x A) = Suc(card(A-{x}))";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   391
by (case_tac "x:A" 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   392
by (asm_simp_tac (!simpset addsimps [insert_absorb]) 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   393
by (dtac mk_disjoint_insert 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   394
by (etac exE 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   395
by (Asm_simp_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   396
by (rtac card_insert_disjoint 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   397
by (rtac (major RSN (2,finite_subset)) 1);
2922
580647a879cf Using Blast_tac
paulson
parents: 2031
diff changeset
   398
by (Blast_tac 1);
580647a879cf Using Blast_tac
paulson
parents: 2031
diff changeset
   399
by (Blast_tac 1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   400
by (asm_simp_tac (!simpset addsimps [major RS card_insert_disjoint]) 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   401
qed "card_insert";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   402
Addsimps [card_insert];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   403
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   404
3340
a886795c9dce Two results suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   405
goal Finite.thy "!!A. finite(A) ==> inj_onto f A --> card (f `` A) = card A";
a886795c9dce Two results suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   406
by (etac finite_induct 1);
a886795c9dce Two results suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   407
by (ALLGOALS Asm_simp_tac);
a886795c9dce Two results suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   408
by (Step_tac 1);
a886795c9dce Two results suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   409
bw inj_onto_def;
a886795c9dce Two results suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   410
by (Blast_tac 1);
a886795c9dce Two results suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   411
by (stac card_insert_disjoint 1);
a886795c9dce Two results suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   412
by (etac finite_imageI 1);
a886795c9dce Two results suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   413
by (Blast_tac 1);
a886795c9dce Two results suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   414
by (Blast_tac 1);
a886795c9dce Two results suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   415
qed_spec_mp "card_image";
a886795c9dce Two results suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   416
a886795c9dce Two results suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   417
3222
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   418
goalw Finite.thy [psubset_def]
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   419
"!!B. finite B ==> !A. A < B --> card(A) < card(B)";
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   420
by (etac finite_induct 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   421
by (Simp_tac 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   422
by (Blast_tac 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   423
by (strip_tac 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   424
by (etac conjE 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   425
by (case_tac "x:A" 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   426
(*1*)
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   427
by (dtac mk_disjoint_insert 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   428
by (etac exE 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   429
by (etac conjE 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   430
by (hyp_subst_tac 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   431
by (rotate_tac ~1 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   432
by (asm_full_simp_tac (!simpset addsimps [subset_insert_iff,finite_subset]) 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   433
by (dtac insert_lim 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   434
by (Asm_full_simp_tac 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   435
(*2*)
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   436
by (rotate_tac ~1 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   437
by (asm_full_simp_tac (!simpset addsimps [subset_insert_iff,finite_subset]) 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   438
by (case_tac "A=F" 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   439
by (Asm_simp_tac 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   440
by (Asm_simp_tac 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   441
qed_spec_mp "psubset_card" ;
3368
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   442
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   443
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   444
(*Relates to equivalence classes.   Based on a theorem of F. Kammüller's.
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   445
  The "finite C" premise is redundant*)
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   446
goal thy "!!C. finite C ==> finite (Union C) --> \
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   447
\          (! c : C. k dvd card c) -->  \
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   448
\          (! c1: C. ! c2: C. c1 ~= c2 --> c1 Int c2 = {}) \
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   449
\          --> k dvd card(Union C)";
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   450
by (etac finite_induct 1);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   451
by (ALLGOALS Asm_simp_tac);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   452
by (strip_tac 1);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   453
by (REPEAT (etac conjE 1));
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   454
by (stac card_Un_disjoint 1);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   455
by (ALLGOALS
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   456
    (asm_full_simp_tac (!simpset
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   457
			 addsimps [dvd_add, disjoint_eq_subset_Compl])));
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   458
by (thin_tac "?PP-->?QQ" 1);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   459
by (thin_tac "!c:F. ?PP(c)" 1);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   460
by (thin_tac "!c:F. ?PP(c) & ?QQ(c)" 1);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   461
by (Step_tac 1);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   462
by (ball_tac 1);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   463
by (Blast_tac 1);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   464
qed_spec_mp "dvd_partition";
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   465