src/HOL/Finite.ML
author paulson
Wed, 21 Aug 1996 13:22:23 +0200
changeset 1933 8b24773de6db
parent 1786 8a31d85d27b8
child 2031 03a843f0f447
permissions -rw-r--r--
Addition of message NS5
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)";
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1660
diff changeset
    19
by (fast_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))";
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1660
diff changeset
    58
by (fast_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);
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1660
diff changeset
    66
by (fast_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]
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
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);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    86
by (rtac (Diff_insert RS ssubst) 2);
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);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   135
qed "subset_finite";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   136
Addsimps[subset_finite];
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);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   140
qed "insert_finite";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   141
Addsimps[insert_finite];
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
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   147
(*The image of a finite set is finite*)
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   148
goal Finite.thy "!!F. finite F ==> finite(h``F)";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   149
by (etac finite_induct 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   150
by (ALLGOALS Asm_simp_tac);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   151
qed "finite_imageI";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   152
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   153
val major::prems = goalw Finite.thy [finite_def]
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   154
    "[| finite A;                                       \
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   155
\       P(A);                                           \
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   156
\       !!a A. [| finite A;  a:A;  P(A) |] ==> P(A-{a}) \
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   157
\    |] ==> P({})";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   158
by (rtac (major RS Fin_empty_induct) 1);
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   159
by (REPEAT (ares_tac (subset_refl::prems) 1));
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   160
qed "finite_empty_induct";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   161
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   162
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   163
section "Finite cardinality -- 'card'";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   164
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   165
goal Set.thy "{f i |i. P i | i=n} = insert (f n) {f i|i. P i}";
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1660
diff changeset
   166
by (Fast_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   167
val Collect_conv_insert = result();
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   168
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   169
goalw Finite.thy [card_def] "card {} = 0";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   170
by (rtac Least_equality 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   171
by (ALLGOALS Asm_full_simp_tac);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   172
qed "card_empty";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   173
Addsimps [card_empty];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   174
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   175
val [major] = goal Finite.thy
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   176
  "finite A ==> ? (n::nat) f. A = {f i |i. i<n}";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   177
by (rtac (major RS finite_induct) 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   178
 by (res_inst_tac [("x","0")] exI 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   179
 by (Simp_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   180
by (etac exE 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   181
by (etac exE 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   182
by (hyp_subst_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   183
by (res_inst_tac [("x","Suc n")] exI 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   184
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
   185
by (asm_simp_tac (!simpset addsimps [Collect_conv_insert, less_Suc_eq]
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   186
                          addcongs [rev_conj_cong]) 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   187
qed "finite_has_card";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   188
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   189
goal Finite.thy
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   190
  "!!A.[| x ~: A; insert x A = {f i|i.i<n} |] ==> \
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   191
\  ? m::nat. m<n & (? g. A = {g i|i.i<m})";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   192
by (res_inst_tac [("n","n")] natE 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   193
 by (hyp_subst_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   194
 by (Asm_full_simp_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   195
by (rename_tac "m" 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   196
by (hyp_subst_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   197
by (case_tac "? a. a:A" 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   198
 by (res_inst_tac [("x","0")] exI 2);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   199
 by (Simp_tac 2);
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1660
diff changeset
   200
 by (Fast_tac 2);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   201
by (etac exE 1);
1660
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   202
by (simp_tac (!simpset addsimps [less_Suc_eq]) 1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   203
by (rtac exI 1);
1782
ab45b881fa62 Shortened a proof
paulson
parents: 1760
diff changeset
   204
by (rtac (refl RS disjI2 RS conjI) 1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   205
by (etac equalityE 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   206
by (asm_full_simp_tac
1660
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   207
     (!simpset addsimps [subset_insert,Collect_conv_insert, less_Suc_eq]) 1);
1786
8a31d85d27b8 best_tac, deepen_tac and safe_tac now also use default claset.
berghofe
parents: 1782
diff changeset
   208
by (SELECT_GOAL(safe_tac (!claset))1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   209
  by (Asm_full_simp_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   210
  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
   211
  by (SELECT_GOAL(safe_tac (!claset))1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   212
   by (subgoal_tac "x ~= f m" 1);
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1660
diff changeset
   213
    by (Fast_tac 2);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   214
   by (subgoal_tac "? k. f k = x & k<m" 1);
1786
8a31d85d27b8 best_tac, deepen_tac and safe_tac now also use default claset.
berghofe
parents: 1782
diff changeset
   215
    by (best_tac (!claset) 2);
8a31d85d27b8 best_tac, deepen_tac and safe_tac now also use default claset.
berghofe
parents: 1782
diff changeset
   216
   by (SELECT_GOAL(safe_tac (!claset))1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   217
   by (res_inst_tac [("x","k")] exI 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   218
   by (Asm_simp_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   219
  by (simp_tac (!simpset setloop (split_tac [expand_if])) 1);
1786
8a31d85d27b8 best_tac, deepen_tac and safe_tac now also use default claset.
berghofe
parents: 1782
diff changeset
   220
  by (best_tac (!claset) 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   221
 bd sym 1;
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   222
 by (rotate_tac ~1 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   223
 by (Asm_full_simp_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   224
 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
   225
 by (SELECT_GOAL(safe_tac (!claset))1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   226
  by (subgoal_tac "x ~= f m" 1);
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1660
diff changeset
   227
   by (Fast_tac 2);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   228
  by (subgoal_tac "? k. f k = x & k<m" 1);
1786
8a31d85d27b8 best_tac, deepen_tac and safe_tac now also use default claset.
berghofe
parents: 1782
diff changeset
   229
   by (best_tac (!claset) 2);
8a31d85d27b8 best_tac, deepen_tac and safe_tac now also use default claset.
berghofe
parents: 1782
diff changeset
   230
  by (SELECT_GOAL(safe_tac (!claset))1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   231
  by (res_inst_tac [("x","k")] exI 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   232
  by (Asm_simp_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   233
 by (simp_tac (!simpset setloop (split_tac [expand_if])) 1);
1786
8a31d85d27b8 best_tac, deepen_tac and safe_tac now also use default claset.
berghofe
parents: 1782
diff changeset
   234
 by (best_tac (!claset) 1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   235
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
   236
by (SELECT_GOAL(safe_tac (!claset))1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   237
 by (subgoal_tac "x ~= f i" 1);
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1660
diff changeset
   238
  by (Fast_tac 2);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   239
 by (case_tac "x = f m" 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   240
  by (res_inst_tac [("x","i")] exI 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   241
  by (Asm_simp_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   242
 by (subgoal_tac "? k. f k = x & k<m" 1);
1786
8a31d85d27b8 best_tac, deepen_tac and safe_tac now also use default claset.
berghofe
parents: 1782
diff changeset
   243
  by (best_tac (!claset) 2);
8a31d85d27b8 best_tac, deepen_tac and safe_tac now also use default claset.
berghofe
parents: 1782
diff changeset
   244
 by (SELECT_GOAL(safe_tac (!claset))1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   245
 by (res_inst_tac [("x","k")] exI 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   246
 by (Asm_simp_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   247
by (simp_tac (!simpset setloop (split_tac [expand_if])) 1);
1786
8a31d85d27b8 best_tac, deepen_tac and safe_tac now also use default claset.
berghofe
parents: 1782
diff changeset
   248
by (best_tac (!claset) 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   249
val lemma = result();
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   250
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   251
goal Finite.thy "!!A. [| finite A; x ~: A |] ==> \
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   252
\ (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
   253
by (rtac Least_equality 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   254
 bd finite_has_card 1;
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   255
 be exE 1;
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   256
 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
   257
 be exE 1;
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   258
 by (res_inst_tac
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   259
   [("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
   260
 by (simp_tac
1660
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   261
    (!simpset addsimps [Collect_conv_insert, less_Suc_eq] 
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   262
	      addcongs [rev_conj_cong]) 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   263
 be subst 1;
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   264
 br refl 1;
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   265
by (rtac notI 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   266
by (etac exE 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   267
by (dtac lemma 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   268
 ba 1;
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   269
by (etac exE 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   270
by (etac conjE 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   271
by (dres_inst_tac [("P","%x. ? g. A = {g i |i. i < x}")] Least_le 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   272
by (dtac le_less_trans 1 THEN atac 1);
1660
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   273
by (asm_full_simp_tac (!simpset addsimps [less_Suc_eq]) 1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   274
by (etac disjE 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   275
by (etac less_asym 1 THEN atac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   276
by (hyp_subst_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   277
by (Asm_full_simp_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   278
val lemma = result();
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   279
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   280
goalw Finite.thy [card_def]
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   281
  "!!A. [| finite A; x ~: A |] ==> card(insert x A) = Suc(card A)";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   282
by (etac lemma 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   283
by (assume_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   284
qed "card_insert_disjoint";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   285
1618
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1553
diff changeset
   286
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
   287
by (res_inst_tac [("t", "A")] (insert_Diff RS subst) 1);
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1553
diff changeset
   288
by (assume_tac 1);
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1553
diff changeset
   289
by (asm_simp_tac (!simpset addsimps [card_insert_disjoint]) 1);
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1553
diff changeset
   290
qed "card_Suc_Diff";
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1553
diff changeset
   291
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1553
diff changeset
   292
goal Finite.thy "!!A. [| finite A; x: A |] ==> card(A-{x}) < card A";
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1553
diff changeset
   293
by (resolve_tac [Suc_less_SucD] 1);
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1553
diff changeset
   294
by (asm_simp_tac (!simpset addsimps [card_Suc_Diff]) 1);
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1553
diff changeset
   295
qed "card_Diff";
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1553
diff changeset
   296
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   297
val [major] = goal Finite.thy
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   298
  "finite A ==> card(insert x A) = Suc(card(A-{x}))";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   299
by (case_tac "x:A" 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   300
by (asm_simp_tac (!simpset addsimps [insert_absorb]) 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   301
by (dtac mk_disjoint_insert 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   302
by (etac exE 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 (rtac card_insert_disjoint 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   305
by (rtac (major RSN (2,finite_subset)) 1);
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1660
diff changeset
   306
by (Fast_tac 1);
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1660
diff changeset
   307
by (Fast_tac 1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   308
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
   309
qed "card_insert";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   310
Addsimps [card_insert];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   311
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   312
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   313
goal Finite.thy  "!!A. finite A ==> !B. B <= A --> card(B) <= card(A)";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   314
by (etac finite_induct 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   315
by (Simp_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   316
by (strip_tac 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   317
by (case_tac "x:B" 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   318
 by (dtac mk_disjoint_insert 1);
1786
8a31d85d27b8 best_tac, deepen_tac and safe_tac now also use default claset.
berghofe
parents: 1782
diff changeset
   319
 by (SELECT_GOAL(safe_tac (!claset))1);
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   320
 by (rotate_tac ~1 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   321
 by (asm_full_simp_tac (!simpset addsimps [subset_insert_iff,finite_subset]) 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   322
by (rotate_tac ~1 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   323
by (asm_full_simp_tac (!simpset addsimps [subset_insert_iff,finite_subset]) 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   324
qed_spec_mp "card_mono";