src/HOL/Finite.ML
author oheimb
Wed, 12 Aug 1998 16:23:25 +0200
changeset 5305 513925de8962
parent 5278 a903b66822e2
child 5316 7a8975451a89
permissions -rw-r--r--
cleanup for Fun.thy: merged Update.{thy|ML} into Fun.{thy|ML} moved o_def from HOL.thy to Fun.thy added Id_def to Fun.thy moved image_compose from Set.ML to Fun.ML moved o_apply and o_assoc from simpdata.ML to Fun.ML moved fun_upd_same and fun_upd_other (from Map.ML) to Fun.ML added fun_upd_twist to Fun.ML
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
3413
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    11
section "finite";
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
(*Discharging ~ x:y entails extra work*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    14
val major::prems = goal Finite.thy 
3413
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    15
    "[| finite F;  P({}); \
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    16
\       !!F x. [| finite F;  x ~: F;  P(F) |] ==> P(insert x F) \
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    17
\    |] ==> P(F)";
3413
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    18
by (rtac (major RS Finites.induct) 1);
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    19
by (excluded_middle_tac "a:A" 2);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    20
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
    21
by (REPEAT (ares_tac prems 1));
3413
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    22
qed "finite_induct";
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    23
4386
b3cff8adc213 Tidied proof of finite_subset_induct
paulson
parents: 4304
diff changeset
    24
val major::subs::prems = goal Finite.thy 
3413
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    25
    "[| finite F;  F <= A; \
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    26
\       P({}); \
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    27
\       !!F a. [| finite F; a:A; a ~: F;  P(F) |] ==> P(insert a F) \
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    28
\    |] ==> P(F)";
4386
b3cff8adc213 Tidied proof of finite_subset_induct
paulson
parents: 4304
diff changeset
    29
by (rtac (subs RS rev_mp) 1);
b3cff8adc213 Tidied proof of finite_subset_induct
paulson
parents: 4304
diff changeset
    30
by (rtac (major RS finite_induct) 1);
b3cff8adc213 Tidied proof of finite_subset_induct
paulson
parents: 4304
diff changeset
    31
by (ALLGOALS (blast_tac (claset() addIs prems)));
3413
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    32
qed "finite_subset_induct";
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    33
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    34
Addsimps Finites.intrs;
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    35
AddSIs Finites.intrs;
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    36
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    37
(*The union of two finite sets is finite*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    38
val major::prems = goal Finite.thy
3413
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    39
    "[| finite F;  finite G |] ==> finite(F Un G)";
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    40
by (rtac (major RS finite_induct) 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4059
diff changeset
    41
by (ALLGOALS (asm_simp_tac (simpset() addsimps prems)));
3413
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    42
qed "finite_UnI";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    43
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    44
(*Every subset of a finite set is finite*)
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
    45
Goal "finite B ==> ALL A. A<=B --> finite A";
4304
af2a2cd9fa51 Tidying
paulson
parents: 4153
diff changeset
    46
by (etac finite_induct 1);
af2a2cd9fa51 Tidying
paulson
parents: 4153
diff changeset
    47
by (Simp_tac 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4059
diff changeset
    48
by (safe_tac (claset() addSDs [subset_insert_iff RS iffD1]));
4304
af2a2cd9fa51 Tidying
paulson
parents: 4153
diff changeset
    49
by (eres_inst_tac [("t","A")] (insert_Diff RS subst) 2);
1264
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 923
diff changeset
    50
by (ALLGOALS Asm_simp_tac);
4304
af2a2cd9fa51 Tidying
paulson
parents: 4153
diff changeset
    51
val lemma = result();
af2a2cd9fa51 Tidying
paulson
parents: 4153
diff changeset
    52
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
    53
Goal "[| A<=B;  finite B |] ==> finite A";
4423
a129b817b58a expandshort;
wenzelm
parents: 4386
diff changeset
    54
by (dtac lemma 1);
4304
af2a2cd9fa51 Tidying
paulson
parents: 4153
diff changeset
    55
by (Blast_tac 1);
3413
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    56
qed "finite_subset";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    57
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
    58
Goal "finite(F Un G) = (finite F & finite G)";
4304
af2a2cd9fa51 Tidying
paulson
parents: 4153
diff changeset
    59
by (blast_tac (claset() 
af2a2cd9fa51 Tidying
paulson
parents: 4153
diff changeset
    60
	         addIs [read_instantiate [("B", "?AA Un ?BB")] finite_subset, 
af2a2cd9fa51 Tidying
paulson
parents: 4153
diff changeset
    61
			finite_UnI]) 1);
3413
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    62
qed "finite_Un";
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    63
AddIffs[finite_Un];
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    64
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
    65
Goal "finite(insert a A) = finite A";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
    66
by (stac insert_is_Un 1);
3413
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    67
by (simp_tac (HOL_ss addsimps [finite_Un]) 1);
3427
e7cef2081106 Removed a few redundant additions of simprules or classical rules
paulson
parents: 3416
diff changeset
    68
by (Blast_tac 1);
3413
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    69
qed "finite_insert";
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    70
Addsimps[finite_insert];
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    71
3413
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    72
(*The image of a finite set is finite *)
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
    73
Goal  "finite F ==> finite(h``F)";
3413
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    74
by (etac finite_induct 1);
1264
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 923
diff changeset
    75
by (Simp_tac 1);
3413
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    76
by (Asm_simp_tac 1);
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    77
qed "finite_imageI";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    78
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    79
val major::prems = goal Finite.thy 
3413
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    80
    "[| finite c;  finite b;                                  \
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
    81
\       P(b);                                                   \
3413
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    82
\       !!x y. [| finite y;  x:y;  P(y) |] ==> P(y-{x}) \
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    83
\    |] ==> c<=b --> P(b-c)";
3413
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    84
by (rtac (major RS finite_induct) 1);
2031
03a843f0f447 Ran expandshort
paulson
parents: 1786
diff changeset
    85
by (stac Diff_insert 2);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    86
by (ALLGOALS (asm_simp_tac
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4059
diff changeset
    87
                (simpset() addsimps (prems@[Diff_subset RS finite_subset]))));
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    88
val lemma = result();
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    89
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    90
val prems = goal Finite.thy 
3413
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    91
    "[| finite A;                                       \
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    92
\       P(A);                                           \
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    93
\       !!a A. [| finite A;  a:A;  P(A) |] ==> P(A-{a}) \
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    94
\    |] ==> P({})";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    95
by (rtac (Diff_cancel RS subst) 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    96
by (rtac (lemma RS mp) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    97
by (REPEAT (ares_tac (subset_refl::prems) 1));
3413
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
    98
qed "finite_empty_induct";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    99
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   100
1618
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1553
diff changeset
   101
(* finite B ==> finite (B - Ba) *)
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1553
diff changeset
   102
bind_thm ("finite_Diff", Diff_subset RS finite_subset);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   103
Addsimps [finite_Diff];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   104
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   105
Goal "finite(A-{a}) = finite(A)";
3368
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   106
by (case_tac "a:A" 1);
3457
a8ab7c64817c Ran expandshort
paulson
parents: 3439
diff changeset
   107
by (rtac (finite_insert RS sym RS trans) 1);
3368
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   108
by (stac insert_Diff 1);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   109
by (ALLGOALS Asm_simp_tac);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   110
qed "finite_Diff_singleton";
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   111
AddIffs [finite_Diff_singleton];
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   112
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4014
diff changeset
   113
(*Lemma for proving finite_imageD*)
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
   114
Goal "finite B ==> !A. f``A = B --> inj_on f A --> finite A";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   115
by (etac finite_induct 1);
3413
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
   116
 by (ALLGOALS Asm_simp_tac);
3708
56facaebf3e3 Changed some proofs to use Clarify_tac
paulson
parents: 3517
diff changeset
   117
by (Clarify_tac 1);
3413
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
   118
by (subgoal_tac "EX y:A. f y = x & F = f``(A-{y})" 1);
3708
56facaebf3e3 Changed some proofs to use Clarify_tac
paulson
parents: 3517
diff changeset
   119
 by (Clarify_tac 1);
4830
bd73675adbed Added a few lemmas.
nipkow
parents: 4775
diff changeset
   120
 by (full_simp_tac (simpset() addsimps [inj_on_def]) 1);
3413
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
   121
 by (Blast_tac 1);
3368
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   122
by (thin_tac "ALL A. ?PP(A)" 1);
3413
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
   123
by (forward_tac [[equalityD2, insertI1] MRS subsetD] 1);
3708
56facaebf3e3 Changed some proofs to use Clarify_tac
paulson
parents: 3517
diff changeset
   124
by (Clarify_tac 1);
3368
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   125
by (res_inst_tac [("x","xa")] bexI 1);
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4014
diff changeset
   126
by (ALLGOALS 
4830
bd73675adbed Added a few lemmas.
nipkow
parents: 4775
diff changeset
   127
    (asm_full_simp_tac (simpset() addsimps [inj_on_image_set_diff])));
3368
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   128
val lemma = result();
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   129
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
   130
Goal "[| finite(f``A);  inj_on f A |] ==> finite A";
3457
a8ab7c64817c Ran expandshort
paulson
parents: 3439
diff changeset
   131
by (dtac lemma 1);
3368
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   132
by (Blast_tac 1);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   133
qed "finite_imageD";
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   134
4014
df6cd80b6387 Added finite_UNION/SigmaI.
nipkow
parents: 3919
diff changeset
   135
(** The finite UNION of finite sets **)
df6cd80b6387 Added finite_UNION/SigmaI.
nipkow
parents: 3919
diff changeset
   136
df6cd80b6387 Added finite_UNION/SigmaI.
nipkow
parents: 3919
diff changeset
   137
val [prem] = goal Finite.thy
df6cd80b6387 Added finite_UNION/SigmaI.
nipkow
parents: 3919
diff changeset
   138
 "finite A ==> (!a:A. finite(B a)) --> finite(UN a:A. B a)";
4153
e534c4c32d54 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4089
diff changeset
   139
by (rtac (prem RS finite_induct) 1);
e534c4c32d54 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4089
diff changeset
   140
by (ALLGOALS Asm_simp_tac);
4014
df6cd80b6387 Added finite_UNION/SigmaI.
nipkow
parents: 3919
diff changeset
   141
bind_thm("finite_UnionI", ballI RSN (2, result() RS mp));
df6cd80b6387 Added finite_UNION/SigmaI.
nipkow
parents: 3919
diff changeset
   142
Addsimps [finite_UnionI];
df6cd80b6387 Added finite_UNION/SigmaI.
nipkow
parents: 3919
diff changeset
   143
df6cd80b6387 Added finite_UNION/SigmaI.
nipkow
parents: 3919
diff changeset
   144
(** Sigma of finite sets **)
df6cd80b6387 Added finite_UNION/SigmaI.
nipkow
parents: 3919
diff changeset
   145
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   146
Goalw [Sigma_def]
5148
74919e8f221c More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5143
diff changeset
   147
 "[| finite A; !a:A. finite(B a) |] ==> finite(SIGMA a:A. B a)";
4153
e534c4c32d54 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4089
diff changeset
   148
by (blast_tac (claset() addSIs [finite_UnionI]) 1);
4014
df6cd80b6387 Added finite_UNION/SigmaI.
nipkow
parents: 3919
diff changeset
   149
bind_thm("finite_SigmaI", ballI RSN (2,result()));
df6cd80b6387 Added finite_UNION/SigmaI.
nipkow
parents: 3919
diff changeset
   150
Addsimps [finite_SigmaI];
3368
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   151
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   152
(** The powerset of a finite set **)
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   153
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
   154
Goal "finite(Pow A) ==> finite A";
3368
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   155
by (subgoal_tac "finite ((%x.{x})``A)" 1);
3457
a8ab7c64817c Ran expandshort
paulson
parents: 3439
diff changeset
   156
by (rtac finite_subset 2);
a8ab7c64817c Ran expandshort
paulson
parents: 3439
diff changeset
   157
by (assume_tac 3);
3368
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   158
by (ALLGOALS
4830
bd73675adbed Added a few lemmas.
nipkow
parents: 4775
diff changeset
   159
    (fast_tac (claset() addSDs [rewrite_rule [inj_on_def] finite_imageD])));
3368
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   160
val lemma = result();
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   161
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   162
Goal "finite(Pow A) = finite A";
3457
a8ab7c64817c Ran expandshort
paulson
parents: 3439
diff changeset
   163
by (rtac iffI 1);
a8ab7c64817c Ran expandshort
paulson
parents: 3439
diff changeset
   164
by (etac lemma 1);
3368
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   165
(*Opposite inclusion: finite A ==> finite (Pow A) *)
3340
a886795c9dce Two results suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   166
by (etac finite_induct 1);
a886795c9dce Two results suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   167
by (ALLGOALS 
a886795c9dce Two results suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   168
    (asm_simp_tac
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4059
diff changeset
   169
     (simpset() addsimps [finite_UnI, finite_imageI, Pow_insert])));
3368
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   170
qed "finite_Pow_iff";
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   171
AddIffs [finite_Pow_iff];
3340
a886795c9dce Two results suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   172
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   173
Goal "finite(r^-1) = finite r";
3457
a8ab7c64817c Ran expandshort
paulson
parents: 3439
diff changeset
   174
by (subgoal_tac "r^-1 = (%(x,y).(y,x))``r" 1);
a8ab7c64817c Ran expandshort
paulson
parents: 3439
diff changeset
   175
 by (Asm_simp_tac 1);
a8ab7c64817c Ran expandshort
paulson
parents: 3439
diff changeset
   176
 by (rtac iffI 1);
4830
bd73675adbed Added a few lemmas.
nipkow
parents: 4775
diff changeset
   177
  by (etac (rewrite_rule [inj_on_def] finite_imageD) 1);
bd73675adbed Added a few lemmas.
nipkow
parents: 4775
diff changeset
   178
  by (simp_tac (simpset() addsplits [split_split]) 1);
3457
a8ab7c64817c Ran expandshort
paulson
parents: 3439
diff changeset
   179
 by (etac finite_imageI 1);
4746
a5dcd7e4a37d inverse -> converse
paulson
parents: 4686
diff changeset
   180
by (simp_tac (simpset() addsimps [converse_def,image_def]) 1);
4477
b3e5857d8d99 New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents: 4423
diff changeset
   181
by Auto_tac;
3457
a8ab7c64817c Ran expandshort
paulson
parents: 3439
diff changeset
   182
 by (rtac bexI 1);
a8ab7c64817c Ran expandshort
paulson
parents: 3439
diff changeset
   183
 by (assume_tac 2);
4763
56072b72d730 adapted proof of finite_converse
oheimb
parents: 4746
diff changeset
   184
by (Simp_tac 1);
4746
a5dcd7e4a37d inverse -> converse
paulson
parents: 4686
diff changeset
   185
qed "finite_converse";
a5dcd7e4a37d inverse -> converse
paulson
parents: 4686
diff changeset
   186
AddIffs [finite_converse];
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   187
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   188
section "Finite cardinality -- 'card'";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   189
4304
af2a2cd9fa51 Tidying
paulson
parents: 4153
diff changeset
   190
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
   191
by (Blast_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   192
val Collect_conv_insert = result();
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   193
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   194
Goalw [card_def] "card {} = 0";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   195
by (rtac Least_equality 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   196
by (ALLGOALS Asm_full_simp_tac);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   197
qed "card_empty";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   198
Addsimps [card_empty];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   199
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   200
val [major] = goal Finite.thy
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   201
  "finite A ==> ? (n::nat) f. A = {f i |i. i<n}";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   202
by (rtac (major RS finite_induct) 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   203
 by (res_inst_tac [("x","0")] exI 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   204
 by (Simp_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   205
by (etac exE 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   206
by (etac exE 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   207
by (hyp_subst_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   208
by (res_inst_tac [("x","Suc n")] exI 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   209
by (res_inst_tac [("x","%i. if i<n then f i else x")] exI 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4059
diff changeset
   210
by (asm_simp_tac (simpset() addsimps [Collect_conv_insert, less_Suc_eq]
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   211
                          addcongs [rev_conj_cong]) 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   212
qed "finite_has_card";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   213
5278
a903b66822e2 even more tidying of Goal commands
paulson
parents: 5183
diff changeset
   214
Goal "[| x ~: A; insert x A = {f i|i. i<n} |]  \
a903b66822e2 even more tidying of Goal commands
paulson
parents: 5183
diff changeset
   215
\     ==> ? m::nat. m<n & (? g. A = {g i|i. i<m})";
5183
89f162de39cf Adapted to new datatype package.
berghofe
parents: 5148
diff changeset
   216
by (exhaust_tac "n" 1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   217
 by (hyp_subst_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   218
 by (Asm_full_simp_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   219
by (rename_tac "m" 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   220
by (hyp_subst_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   221
by (case_tac "? a. a:A" 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   222
 by (res_inst_tac [("x","0")] exI 2);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   223
 by (Simp_tac 2);
2922
580647a879cf Using Blast_tac
paulson
parents: 2031
diff changeset
   224
 by (Blast_tac 2);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   225
by (etac exE 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4059
diff changeset
   226
by (simp_tac (simpset() addsimps [less_Suc_eq]) 1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   227
by (rtac exI 1);
1782
ab45b881fa62 Shortened a proof
paulson
parents: 1760
diff changeset
   228
by (rtac (refl RS disjI2 RS conjI) 1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   229
by (etac equalityE 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   230
by (asm_full_simp_tac
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4059
diff changeset
   231
     (simpset() addsimps [subset_insert,Collect_conv_insert, less_Suc_eq]) 1);
4153
e534c4c32d54 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4089
diff changeset
   232
by Safe_tac;
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   233
  by (Asm_full_simp_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   234
  by (res_inst_tac [("x","%i. if f i = f m then a else f i")] exI 1);
4153
e534c4c32d54 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4089
diff changeset
   235
  by (SELECT_GOAL Safe_tac 1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   236
   by (subgoal_tac "x ~= f m" 1);
2922
580647a879cf Using Blast_tac
paulson
parents: 2031
diff changeset
   237
    by (Blast_tac 2);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   238
   by (subgoal_tac "? k. f k = x & k<m" 1);
2922
580647a879cf Using Blast_tac
paulson
parents: 2031
diff changeset
   239
    by (Blast_tac 2);
4153
e534c4c32d54 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4089
diff changeset
   240
   by (SELECT_GOAL Safe_tac 1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   241
   by (res_inst_tac [("x","k")] exI 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   242
   by (Asm_simp_tac 1);
4686
74a12e86b20b Removed `addsplits [expand_if]'
nipkow
parents: 4477
diff changeset
   243
  by (Simp_tac 1);
2922
580647a879cf Using Blast_tac
paulson
parents: 2031
diff changeset
   244
  by (Blast_tac 1);
3457
a8ab7c64817c Ran expandshort
paulson
parents: 3439
diff changeset
   245
 by (dtac sym 1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   246
 by (rotate_tac ~1 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   247
 by (Asm_full_simp_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   248
 by (res_inst_tac [("x","%i. if f i = f m then a else f i")] exI 1);
4153
e534c4c32d54 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4089
diff changeset
   249
 by (SELECT_GOAL Safe_tac 1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   250
  by (subgoal_tac "x ~= f m" 1);
2922
580647a879cf Using Blast_tac
paulson
parents: 2031
diff changeset
   251
   by (Blast_tac 2);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   252
  by (subgoal_tac "? k. f k = x & k<m" 1);
2922
580647a879cf Using Blast_tac
paulson
parents: 2031
diff changeset
   253
   by (Blast_tac 2);
4153
e534c4c32d54 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4089
diff changeset
   254
  by (SELECT_GOAL Safe_tac 1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   255
  by (res_inst_tac [("x","k")] exI 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   256
  by (Asm_simp_tac 1);
4686
74a12e86b20b Removed `addsplits [expand_if]'
nipkow
parents: 4477
diff changeset
   257
 by (Simp_tac 1);
2922
580647a879cf Using Blast_tac
paulson
parents: 2031
diff changeset
   258
 by (Blast_tac 1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   259
by (res_inst_tac [("x","%j. if f j = f i then f m else f j")] exI 1);
4153
e534c4c32d54 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4089
diff changeset
   260
by (SELECT_GOAL Safe_tac 1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   261
 by (subgoal_tac "x ~= f i" 1);
2922
580647a879cf Using Blast_tac
paulson
parents: 2031
diff changeset
   262
  by (Blast_tac 2);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   263
 by (case_tac "x = f m" 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   264
  by (res_inst_tac [("x","i")] exI 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   265
  by (Asm_simp_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   266
 by (subgoal_tac "? k. f k = x & k<m" 1);
2922
580647a879cf Using Blast_tac
paulson
parents: 2031
diff changeset
   267
  by (Blast_tac 2);
4153
e534c4c32d54 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4089
diff changeset
   268
 by (SELECT_GOAL Safe_tac 1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   269
 by (res_inst_tac [("x","k")] exI 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   270
 by (Asm_simp_tac 1);
4686
74a12e86b20b Removed `addsplits [expand_if]'
nipkow
parents: 4477
diff changeset
   271
by (Simp_tac 1);
2922
580647a879cf Using Blast_tac
paulson
parents: 2031
diff changeset
   272
by (Blast_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   273
val lemma = result();
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   274
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
   275
Goal "[| finite A; x ~: A |] ==> \
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3724
diff changeset
   276
\ (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
   277
by (rtac Least_equality 1);
3457
a8ab7c64817c Ran expandshort
paulson
parents: 3439
diff changeset
   278
 by (dtac finite_has_card 1);
a8ab7c64817c Ran expandshort
paulson
parents: 3439
diff changeset
   279
 by (etac exE 1);
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3724
diff changeset
   280
 by (dres_inst_tac [("P","%n.? f. A={f i|i. i<n}")] LeastI 1);
3457
a8ab7c64817c Ran expandshort
paulson
parents: 3439
diff changeset
   281
 by (etac exE 1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   282
 by (res_inst_tac
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   283
   [("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
   284
 by (simp_tac
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4059
diff changeset
   285
    (simpset() addsimps [Collect_conv_insert, less_Suc_eq] 
2031
03a843f0f447 Ran expandshort
paulson
parents: 1786
diff changeset
   286
              addcongs [rev_conj_cong]) 1);
3457
a8ab7c64817c Ran expandshort
paulson
parents: 3439
diff changeset
   287
 by (etac subst 1);
a8ab7c64817c Ran expandshort
paulson
parents: 3439
diff changeset
   288
 by (rtac refl 1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   289
by (rtac notI 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   290
by (etac exE 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   291
by (dtac lemma 1);
3457
a8ab7c64817c Ran expandshort
paulson
parents: 3439
diff changeset
   292
 by (assume_tac 1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   293
by (etac exE 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   294
by (etac conjE 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   295
by (dres_inst_tac [("P","%x. ? g. A = {g i |i. i < x}")] Least_le 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   296
by (dtac le_less_trans 1 THEN atac 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4059
diff changeset
   297
by (asm_full_simp_tac (simpset() addsimps [less_Suc_eq]) 1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   298
by (etac disjE 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   299
by (etac less_asym 1 THEN atac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   300
by (hyp_subst_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   301
by (Asm_full_simp_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   302
val lemma = result();
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   303
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   304
Goalw [card_def]
5148
74919e8f221c More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5143
diff changeset
   305
  "[| finite A; x ~: A |] ==> card(insert x A) = Suc(card A)";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   306
by (etac lemma 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   307
by (assume_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   308
qed "card_insert_disjoint";
3352
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   309
Addsimps [card_insert_disjoint];
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   310
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
   311
Goal "finite A ==> card A <= card (insert x A)";
4768
c342d63173e9 New theorems card_Diff_le and card_insert_le; tidied
paulson
parents: 4763
diff changeset
   312
by (case_tac "x: A" 1);
c342d63173e9 New theorems card_Diff_le and card_insert_le; tidied
paulson
parents: 4763
diff changeset
   313
by (ALLGOALS (asm_simp_tac (simpset() addsimps [insert_absorb])));
c342d63173e9 New theorems card_Diff_le and card_insert_le; tidied
paulson
parents: 4763
diff changeset
   314
qed "card_insert_le";
c342d63173e9 New theorems card_Diff_le and card_insert_le; tidied
paulson
parents: 4763
diff changeset
   315
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
   316
Goal  "finite A ==> !B. B <= A --> card(B) <= card(A)";
3352
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   317
by (etac finite_induct 1);
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   318
by (Simp_tac 1);
3708
56facaebf3e3 Changed some proofs to use Clarify_tac
paulson
parents: 3517
diff changeset
   319
by (Clarify_tac 1);
3352
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   320
by (case_tac "x:B" 1);
3413
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
   321
 by (dres_inst_tac [("A","B")] mk_disjoint_insert 1);
4775
66b1a7c42d94 Tidied proofs
paulson
parents: 4768
diff changeset
   322
by (asm_full_simp_tac (simpset() addsimps [subset_insert_iff]) 2);
66b1a7c42d94 Tidied proofs
paulson
parents: 4768
diff changeset
   323
by (fast_tac (claset() addss
66b1a7c42d94 Tidied proofs
paulson
parents: 4768
diff changeset
   324
	      (simpset() addsimps [subset_insert_iff, finite_subset])) 1);
3352
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   325
qed_spec_mp "card_mono";
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   326
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
   327
Goal "[| finite A; finite B |]\
3352
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   328
\                       ==> A Int B = {} --> card(A Un B) = card A + card B";
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   329
by (etac finite_induct 1);
4686
74a12e86b20b Removed `addsplits [expand_if]'
nipkow
parents: 4477
diff changeset
   330
by (ALLGOALS (asm_simp_tac (simpset() addsimps [Int_insert_left])));
3352
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   331
qed_spec_mp "card_Un_disjoint";
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   332
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
   333
Goal "[| finite A; B<=A |] ==> card A - card B = card (A - B)";
3352
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   334
by (subgoal_tac "(A-B) Un B = A" 1);
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   335
by (Blast_tac 2);
3457
a8ab7c64817c Ran expandshort
paulson
parents: 3439
diff changeset
   336
by (rtac (add_right_cancel RS iffD1) 1);
a8ab7c64817c Ran expandshort
paulson
parents: 3439
diff changeset
   337
by (rtac (card_Un_disjoint RS subst) 1);
a8ab7c64817c Ran expandshort
paulson
parents: 3439
diff changeset
   338
by (etac ssubst 4);
3352
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   339
by (Blast_tac 3);
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   340
by (ALLGOALS 
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   341
    (asm_simp_tac
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4059
diff changeset
   342
     (simpset() addsimps [add_commute, not_less_iff_le, 
3352
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   343
			 add_diff_inverse, card_mono, finite_subset])));
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   344
qed "card_Diff_subset";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   345
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
   346
Goal "[| finite A; x: A |] ==> Suc(card(A-{x})) = card A";
1618
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1553
diff changeset
   347
by (res_inst_tac [("t", "A")] (insert_Diff RS subst) 1);
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1553
diff changeset
   348
by (assume_tac 1);
3352
04502e5431fb New theorems suggested by Florian Kammueller
paulson
parents: 3340
diff changeset
   349
by (Asm_simp_tac 1);
1618
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1553
diff changeset
   350
qed "card_Suc_Diff";
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1553
diff changeset
   351
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
   352
Goal "[| finite A; x: A |] ==> card(A-{x}) < card A";
2031
03a843f0f447 Ran expandshort
paulson
parents: 1786
diff changeset
   353
by (rtac Suc_less_SucD 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4059
diff changeset
   354
by (asm_simp_tac (simpset() addsimps [card_Suc_Diff]) 1);
1618
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1553
diff changeset
   355
qed "card_Diff";
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1553
diff changeset
   356
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
   357
Goal "finite A ==> card(A-{x}) <= card A";
4768
c342d63173e9 New theorems card_Diff_le and card_insert_le; tidied
paulson
parents: 4763
diff changeset
   358
by (case_tac "x: A" 1);
c342d63173e9 New theorems card_Diff_le and card_insert_le; tidied
paulson
parents: 4763
diff changeset
   359
by (ALLGOALS (asm_simp_tac (simpset() addsimps [card_Diff, less_imp_le])));
c342d63173e9 New theorems card_Diff_le and card_insert_le; tidied
paulson
parents: 4763
diff changeset
   360
qed "card_Diff_le";
c342d63173e9 New theorems card_Diff_le and card_insert_le; tidied
paulson
parents: 4763
diff changeset
   361
3389
3150eba724a1 New theorem about the cardinality of the powerset (uses exponentiation)
paulson
parents: 3382
diff changeset
   362
3150eba724a1 New theorem about the cardinality of the powerset (uses exponentiation)
paulson
parents: 3382
diff changeset
   363
(*** Cardinality of the Powerset ***)
3150eba724a1 New theorem about the cardinality of the powerset (uses exponentiation)
paulson
parents: 3382
diff changeset
   364
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
   365
Goal "finite A ==> card(insert x A) = Suc(card(A-{x}))";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   366
by (case_tac "x:A" 1);
4768
c342d63173e9 New theorems card_Diff_le and card_insert_le; tidied
paulson
parents: 4763
diff changeset
   367
by (ALLGOALS 
c342d63173e9 New theorems card_Diff_le and card_insert_le; tidied
paulson
parents: 4763
diff changeset
   368
    (asm_simp_tac (simpset() addsimps [card_Suc_Diff, insert_absorb])));
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   369
qed "card_insert";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   370
Addsimps [card_insert];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   371
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
   372
Goal "finite(A) ==> inj_on f A --> card (f `` A) = card A";
3340
a886795c9dce Two results suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   373
by (etac finite_induct 1);
a886795c9dce Two results suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   374
by (ALLGOALS Asm_simp_tac);
3724
f33e301a89f5 Step_tac -> Safe_tac
paulson
parents: 3708
diff changeset
   375
by Safe_tac;
4830
bd73675adbed Added a few lemmas.
nipkow
parents: 4775
diff changeset
   376
by (rewtac inj_on_def);
3340
a886795c9dce Two results suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   377
by (Blast_tac 1);
a886795c9dce Two results suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   378
by (stac card_insert_disjoint 1);
a886795c9dce Two results suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   379
by (etac finite_imageI 1);
a886795c9dce Two results suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   380
by (Blast_tac 1);
a886795c9dce Two results suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   381
by (Blast_tac 1);
a886795c9dce Two results suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   382
qed_spec_mp "card_image";
a886795c9dce Two results suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   383
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
   384
Goal "finite A ==> card (Pow A) = 2 ^ card A";
3389
3150eba724a1 New theorem about the cardinality of the powerset (uses exponentiation)
paulson
parents: 3382
diff changeset
   385
by (etac finite_induct 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4059
diff changeset
   386
by (ALLGOALS (asm_simp_tac (simpset() addsimps [Pow_insert])));
3389
3150eba724a1 New theorem about the cardinality of the powerset (uses exponentiation)
paulson
parents: 3382
diff changeset
   387
by (stac card_Un_disjoint 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4059
diff changeset
   388
by (EVERY (map (blast_tac (claset() addIs [finite_imageI])) [3,2,1]));
4830
bd73675adbed Added a few lemmas.
nipkow
parents: 4775
diff changeset
   389
by (subgoal_tac "inj_on (insert x) (Pow F)" 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4059
diff changeset
   390
by (asm_simp_tac (simpset() addsimps [card_image, Pow_insert]) 1);
4830
bd73675adbed Added a few lemmas.
nipkow
parents: 4775
diff changeset
   391
by (rewtac inj_on_def);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4059
diff changeset
   392
by (blast_tac (claset() addSEs [equalityE]) 1);
3389
3150eba724a1 New theorem about the cardinality of the powerset (uses exponentiation)
paulson
parents: 3382
diff changeset
   393
qed "card_Pow";
3150eba724a1 New theorem about the cardinality of the powerset (uses exponentiation)
paulson
parents: 3382
diff changeset
   394
Addsimps [card_Pow];
3340
a886795c9dce Two results suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   395
3389
3150eba724a1 New theorem about the cardinality of the powerset (uses exponentiation)
paulson
parents: 3382
diff changeset
   396
3150eba724a1 New theorem about the cardinality of the powerset (uses exponentiation)
paulson
parents: 3382
diff changeset
   397
(*Proper subsets*)
5148
74919e8f221c More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5143
diff changeset
   398
Goalw [psubset_def] "finite B ==> !A. A < B --> card(A) < card(B)";
3222
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   399
by (etac finite_induct 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   400
by (Simp_tac 1);
3708
56facaebf3e3 Changed some proofs to use Clarify_tac
paulson
parents: 3517
diff changeset
   401
by (Clarify_tac 1);
3222
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   402
by (case_tac "x:A" 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   403
(*1*)
3413
c1f63cc3a768 Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents: 3389
diff changeset
   404
by (dres_inst_tac [("A","A")]mk_disjoint_insert 1);
4775
66b1a7c42d94 Tidied proofs
paulson
parents: 4768
diff changeset
   405
by (Clarify_tac 1);
66b1a7c42d94 Tidied proofs
paulson
parents: 4768
diff changeset
   406
by (rotate_tac ~3 1);
66b1a7c42d94 Tidied proofs
paulson
parents: 4768
diff changeset
   407
by (asm_full_simp_tac (simpset() addsimps [finite_subset]) 1);
3708
56facaebf3e3 Changed some proofs to use Clarify_tac
paulson
parents: 3517
diff changeset
   408
by (Blast_tac 1);
3222
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   409
(*2*)
3708
56facaebf3e3 Changed some proofs to use Clarify_tac
paulson
parents: 3517
diff changeset
   410
by (eres_inst_tac [("P","?a<?b")] notE 1);
4775
66b1a7c42d94 Tidied proofs
paulson
parents: 4768
diff changeset
   411
by (asm_full_simp_tac (simpset() addsimps [subset_insert_iff]) 1);
3222
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   412
by (case_tac "A=F" 1);
3708
56facaebf3e3 Changed some proofs to use Clarify_tac
paulson
parents: 3517
diff changeset
   413
by (ALLGOALS Asm_simp_tac);
3222
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   414
qed_spec_mp "psubset_card" ;
3368
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   415
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   416
3430
d21b920363ab eliminated non-ASCII;
wenzelm
parents: 3427
diff changeset
   417
(*Relates to equivalence classes.   Based on a theorem of F. Kammueller's.
3368
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   418
  The "finite C" premise is redundant*)
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
   419
Goal "finite C ==> finite (Union C) --> \
3368
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   420
\          (! c : C. k dvd card c) -->  \
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   421
\          (! c1: C. ! c2: C. c1 ~= c2 --> c1 Int c2 = {}) \
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   422
\          --> k dvd card(Union C)";
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   423
by (etac finite_induct 1);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   424
by (ALLGOALS Asm_simp_tac);
3708
56facaebf3e3 Changed some proofs to use Clarify_tac
paulson
parents: 3517
diff changeset
   425
by (Clarify_tac 1);
3368
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   426
by (stac card_Un_disjoint 1);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   427
by (ALLGOALS
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4059
diff changeset
   428
    (asm_full_simp_tac (simpset()
3368
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   429
			 addsimps [dvd_add, disjoint_eq_subset_Compl])));
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   430
by (thin_tac "!c:F. ?PP(c)" 1);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   431
by (thin_tac "!c:F. ?PP(c) & ?QQ(c)" 1);
3708
56facaebf3e3 Changed some proofs to use Clarify_tac
paulson
parents: 3517
diff changeset
   432
by (Clarify_tac 1);
3368
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   433
by (ball_tac 1);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   434
by (Blast_tac 1);
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   435
qed_spec_mp "dvd_partition";
be517d000c02 Many new theorems about cardinality
paulson
parents: 3352
diff changeset
   436