src/ZF/Finite.ML
author oheimb
Wed, 03 Apr 2002 10:21:13 +0200
changeset 13076 70704dd48bd5
parent 9907 473a6604da94
permissions -rw-r--r--
bugfix concerning claset(), added limited support for ALLGOALS + fast_tac etc.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 803
diff changeset
     1
(*  Title:      ZF/Finite.ML
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 803
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
     4
    Copyright   1994  University of Cambridge
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
     5
534
cd8bec47e175 ZF/Finite: added the finite function space, A-||>B
lcp
parents: 516
diff changeset
     6
Finite powerset operator; finite function space
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
     7
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
     8
prove X:Fin(A) ==> |X| < nat
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
     9
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    10
prove:  b: Fin(A) ==> inj(b,b)<=surj(b,b)
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    11
*)
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    12
534
cd8bec47e175 ZF/Finite: added the finite function space, A-||>B
lcp
parents: 516
diff changeset
    13
(*** Finite powerset operator ***)
cd8bec47e175 ZF/Finite: added the finite function space, A-||>B
lcp
parents: 516
diff changeset
    14
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
    15
Goalw Fin.defs "A<=B ==> Fin(A) <= Fin(B)";
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    16
by (rtac lfp_mono 1);
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    17
by (REPEAT (rtac Fin.bnd_mono 1));
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    18
by (REPEAT (ares_tac (Pow_mono::basic_monos) 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 534
diff changeset
    19
qed "Fin_mono";
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    20
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    21
(* A : Fin(B) ==> A <= B *)
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9211
diff changeset
    22
bind_thm ("FinD", Fin.dom_subset RS subsetD RS PowD);
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    23
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    24
(** Induction on finite sets **)
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    25
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    26
(*Discharging x~:y entails extra work*)
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5137
diff changeset
    27
val major::prems = Goal
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    28
    "[| b: Fin(A);  \
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    29
\       P(0);        \
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    30
\       !!x y. [| x: A;  y: Fin(A);  x~:y;  P(y) |] ==> P(cons(x,y)) \
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    31
\    |] ==> P(b)";
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    32
by (rtac (major RS Fin.induct) 1);
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    33
by (excluded_middle_tac "a:b" 2);
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 803
diff changeset
    34
by (etac (cons_absorb RS ssubst) 3 THEN assume_tac 3);      (*backtracking!*)
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    35
by (REPEAT (ares_tac prems 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 534
diff changeset
    36
qed "Fin_induct";
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    37
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    38
(** Simplification for Fin **)
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
    39
Addsimps Fin.intrs;
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    40
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    41
(*The union of two finite sets is finite.*)
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5137
diff changeset
    42
Goal "[| b: Fin(A);  c: Fin(A) |] ==> b Un c : Fin(A)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
    43
by (etac Fin_induct 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
    44
by (ALLGOALS (asm_simp_tac (simpset() addsimps [Un_cons])));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 534
diff changeset
    45
qed "Fin_UnI";
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    46
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
    47
Addsimps [Fin_UnI];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
    48
5412
0c2472c74c24 New lemmas involving Int
paulson
parents: 5268
diff changeset
    49
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    50
(*The union of a set of finite sets is finite.*)
9211
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9173
diff changeset
    51
Goal "C : Fin(Fin(A)) ==> Union(C) : Fin(A)";
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9173
diff changeset
    52
by (etac Fin_induct 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
    53
by (ALLGOALS Asm_simp_tac);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 534
diff changeset
    54
qed "Fin_UnionI";
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    55
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    56
(*Every subset of a finite set is finite.*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
    57
Goal "b: Fin(A) ==> ALL z. z<=b --> z: Fin(A)";
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    58
by (etac Fin_induct 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
    59
by (simp_tac (simpset() addsimps [subset_empty_iff]) 1);
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
    60
by (asm_simp_tac (simpset() addsimps subset_cons_iff::distrib_simps) 1);
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
    61
by Safe_tac;
534
cd8bec47e175 ZF/Finite: added the finite function space, A-||>B
lcp
parents: 516
diff changeset
    62
by (eres_inst_tac [("b","z")] (cons_Diff RS subst) 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
    63
by (Asm_simp_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 534
diff changeset
    64
qed "Fin_subset_lemma";
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    65
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
    66
Goal "[| c<=b;  b: Fin(A) |] ==> c: Fin(A)";
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    67
by (REPEAT (ares_tac [Fin_subset_lemma RS spec RS mp] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 534
diff changeset
    68
qed "Fin_subset";
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    69
5412
0c2472c74c24 New lemmas involving Int
paulson
parents: 5268
diff changeset
    70
Goal "b: Fin(A) ==> b Int c : Fin(A)";
0c2472c74c24 New lemmas involving Int
paulson
parents: 5268
diff changeset
    71
by (blast_tac (claset() addIs [Fin_subset]) 1);
0c2472c74c24 New lemmas involving Int
paulson
parents: 5268
diff changeset
    72
qed "Fin_IntI1";
0c2472c74c24 New lemmas involving Int
paulson
parents: 5268
diff changeset
    73
0c2472c74c24 New lemmas involving Int
paulson
parents: 5268
diff changeset
    74
Goal "c: Fin(A) ==> b Int c : Fin(A)";
0c2472c74c24 New lemmas involving Int
paulson
parents: 5268
diff changeset
    75
by (blast_tac (claset() addIs [Fin_subset]) 1);
0c2472c74c24 New lemmas involving Int
paulson
parents: 5268
diff changeset
    76
qed "Fin_IntI2";
0c2472c74c24 New lemmas involving Int
paulson
parents: 5268
diff changeset
    77
0c2472c74c24 New lemmas involving Int
paulson
parents: 5268
diff changeset
    78
Addsimps[Fin_IntI1, Fin_IntI2];
0c2472c74c24 New lemmas involving Int
paulson
parents: 5268
diff changeset
    79
AddIs[Fin_IntI1, Fin_IntI2];
0c2472c74c24 New lemmas involving Int
paulson
parents: 5268
diff changeset
    80
0c2472c74c24 New lemmas involving Int
paulson
parents: 5268
diff changeset
    81
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5137
diff changeset
    82
val major::prems = Goal
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 803
diff changeset
    83
    "[| c: Fin(A);  b: Fin(A);                                  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 803
diff changeset
    84
\       P(b);                                                   \
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    85
\       !!x y. [| x: A;  y: Fin(A);  x:y;  P(y) |] ==> P(y-{x}) \
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    86
\    |] ==> c<=b --> P(b-c)";
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    87
by (rtac (major RS Fin_induct) 1);
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1956
diff changeset
    88
by (stac Diff_cons 2);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
    89
by (ALLGOALS (asm_simp_tac (simpset() addsimps (prems@[cons_subset_iff, 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 803
diff changeset
    90
                                Diff_subset RS Fin_subset]))));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 534
diff changeset
    91
qed "Fin_0_induct_lemma";
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    92
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5137
diff changeset
    93
val prems = Goal
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 803
diff changeset
    94
    "[| b: Fin(A);                                              \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 803
diff changeset
    95
\       P(b);                                                   \
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    96
\       !!x y. [| x: A;  y: Fin(A);  x:y;  P(y) |] ==> P(y-{x}) \
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    97
\    |] ==> P(0)";
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    98
by (rtac (Diff_cancel RS subst) 1);
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
    99
by (rtac (Fin_0_induct_lemma RS mp) 1);
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
   100
by (REPEAT (ares_tac (subset_refl::prems) 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 534
diff changeset
   101
qed "Fin_0_induct";
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
   102
1957113f0d7d installation of new inductive/datatype sections
lcp
parents:
diff changeset
   103
(*Functions from a finite ordinal*)
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5137
diff changeset
   104
Goal "n: nat ==> n->A <= Fin(nat*A)";
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 5412
diff changeset
   105
by (induct_tac "n" 1);
8201
a81d18b0a9b1 tidied some proofs
paulson
parents: 8183
diff changeset
   106
by (simp_tac (simpset() addsimps [subset_iff]) 1);
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5137
diff changeset
   107
by (asm_simp_tac 
59ef39008514 even more tidying of Goal commands
paulson
parents: 5137
diff changeset
   108
    (simpset() addsimps [succ_def, mem_not_refl RS cons_fun_eq]) 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
   109
by (fast_tac (claset() addSIs [Fin.consI]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 534
diff changeset
   110
qed "nat_fun_subset_Fin";
534
cd8bec47e175 ZF/Finite: added the finite function space, A-||>B
lcp
parents: 516
diff changeset
   111
cd8bec47e175 ZF/Finite: added the finite function space, A-||>B
lcp
parents: 516
diff changeset
   112
cd8bec47e175 ZF/Finite: added the finite function space, A-||>B
lcp
parents: 516
diff changeset
   113
(*** Finite function space ***)
cd8bec47e175 ZF/Finite: added the finite function space, A-||>B
lcp
parents: 516
diff changeset
   114
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   115
Goalw FiniteFun.defs
8181
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   116
    "[| A<=C;  B<=D |] ==> A -||> B  <=  C -||> D";
534
cd8bec47e175 ZF/Finite: added the finite function space, A-||>B
lcp
parents: 516
diff changeset
   117
by (rtac lfp_mono 1);
cd8bec47e175 ZF/Finite: added the finite function space, A-||>B
lcp
parents: 516
diff changeset
   118
by (REPEAT (rtac FiniteFun.bnd_mono 1));
cd8bec47e175 ZF/Finite: added the finite function space, A-||>B
lcp
parents: 516
diff changeset
   119
by (REPEAT (ares_tac (Fin_mono::Sigma_mono::basic_monos) 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 534
diff changeset
   120
qed "FiniteFun_mono";
534
cd8bec47e175 ZF/Finite: added the finite function space, A-||>B
lcp
parents: 516
diff changeset
   121
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   122
Goal "A<=B ==> A -||> A  <=  B -||> B";
534
cd8bec47e175 ZF/Finite: added the finite function space, A-||>B
lcp
parents: 516
diff changeset
   123
by (REPEAT (ares_tac [FiniteFun_mono] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 534
diff changeset
   124
qed "FiniteFun_mono1";
534
cd8bec47e175 ZF/Finite: added the finite function space, A-||>B
lcp
parents: 516
diff changeset
   125
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   126
Goal "h: A -||>B ==> h: domain(h) -> B";
534
cd8bec47e175 ZF/Finite: added the finite function space, A-||>B
lcp
parents: 516
diff changeset
   127
by (etac FiniteFun.induct 1);
8201
a81d18b0a9b1 tidied some proofs
paulson
parents: 8183
diff changeset
   128
by (Simp_tac 1);
a81d18b0a9b1 tidied some proofs
paulson
parents: 8183
diff changeset
   129
by (asm_simp_tac (simpset() addsimps [fun_extend3]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 534
diff changeset
   130
qed "FiniteFun_is_fun";
534
cd8bec47e175 ZF/Finite: added the finite function space, A-||>B
lcp
parents: 516
diff changeset
   131
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   132
Goal "h: A -||>B ==> domain(h) : Fin(A)";
534
cd8bec47e175 ZF/Finite: added the finite function space, A-||>B
lcp
parents: 516
diff changeset
   133
by (etac FiniteFun.induct 1);
8201
a81d18b0a9b1 tidied some proofs
paulson
parents: 8183
diff changeset
   134
by (Simp_tac 1);
a81d18b0a9b1 tidied some proofs
paulson
parents: 8183
diff changeset
   135
by (Asm_simp_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 534
diff changeset
   136
qed "FiniteFun_domain_Fin";
534
cd8bec47e175 ZF/Finite: added the finite function space, A-||>B
lcp
parents: 516
diff changeset
   137
803
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 760
diff changeset
   138
bind_thm ("FiniteFun_apply_type", FiniteFun_is_fun RS apply_type);
534
cd8bec47e175 ZF/Finite: added the finite function space, A-||>B
lcp
parents: 516
diff changeset
   139
cd8bec47e175 ZF/Finite: added the finite function space, A-||>B
lcp
parents: 516
diff changeset
   140
(*Every subset of a finite function is a finite function.*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   141
Goal "b: A-||>B ==> ALL z. z<=b --> z: A-||>B";
534
cd8bec47e175 ZF/Finite: added the finite function space, A-||>B
lcp
parents: 516
diff changeset
   142
by (etac FiniteFun.induct 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
   143
by (simp_tac (simpset() addsimps subset_empty_iff::FiniteFun.intrs) 1);
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
   144
by (asm_simp_tac (simpset() addsimps subset_cons_iff::distrib_simps) 1);
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
   145
by Safe_tac;
534
cd8bec47e175 ZF/Finite: added the finite function space, A-||>B
lcp
parents: 516
diff changeset
   146
by (eres_inst_tac [("b","z")] (cons_Diff RS subst) 1);
cd8bec47e175 ZF/Finite: added the finite function space, A-||>B
lcp
parents: 516
diff changeset
   147
by (dtac (spec RS mp) 1 THEN assume_tac 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
   148
by (fast_tac (claset() addSIs FiniteFun.intrs) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 534
diff changeset
   149
qed "FiniteFun_subset_lemma";
534
cd8bec47e175 ZF/Finite: added the finite function space, A-||>B
lcp
parents: 516
diff changeset
   150
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   151
Goal "[| c<=b;  b: A-||>B |] ==> c: A-||>B";
534
cd8bec47e175 ZF/Finite: added the finite function space, A-||>B
lcp
parents: 516
diff changeset
   152
by (REPEAT (ares_tac [FiniteFun_subset_lemma RS spec RS mp] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 534
diff changeset
   153
qed "FiniteFun_subset";
534
cd8bec47e175 ZF/Finite: added the finite function space, A-||>B
lcp
parents: 516
diff changeset
   154
8181
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   155
(** Some further results by Sidi O. Ehmety **)
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   156
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   157
Goal "A:Fin(X) ==> ALL f. f:A->B --> f:A-||>B";
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   158
by (etac Fin.induct 1);
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   159
by (simp_tac (simpset() addsimps FiniteFun.intrs) 1);
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   160
by (Clarify_tac 1);
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   161
by (case_tac "a:b" 1);
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   162
by (rotate_tac ~1 1);
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   163
by (asm_full_simp_tac (simpset() addsimps [cons_absorb]) 1);
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   164
by (subgoal_tac "restrict(f,b) : b -||> B" 1);
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   165
by (blast_tac (claset() addIs [restrict_type2]) 2);
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   166
by (stac fun_cons_restrict_eq 1 THEN assume_tac 1);
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   167
by (full_simp_tac (simpset() addsimps [restrict_def, lam_def]) 1);
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   168
by (blast_tac (claset() addIs [apply_funtype, impOfSubs FiniteFun_mono]
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   169
                              @FiniteFun.intrs) 1);
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   170
qed_spec_mp "fun_FiniteFunI";
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   171
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   172
Goal "A: Fin(X) ==> (lam x:A. b(x)) : A -||> {b(x). x:A}";
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   173
by (blast_tac (claset() addIs [fun_FiniteFunI, lam_funtype]) 1);
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   174
qed "lam_FiniteFun";
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   175
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   176
Goal "f : FiniteFun(A, {y:B. P(y)})  \
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   177
\     <-> f : FiniteFun(A,B) & (ALL x:domain(f). P(f`x))";
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   178
by Auto_tac;
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   179
by (blast_tac (claset() addIs [impOfSubs FiniteFun_mono]) 1);
9173
422968aeed49 fixed some weak elim rules
paulson
parents: 8201
diff changeset
   180
by (blast_tac (claset() addDs [Pair_mem_PiD, FiniteFun_is_fun]) 1);
8181
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   181
by (res_inst_tac [("A1", "domain(f)")]
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   182
    (subset_refl RSN(2, FiniteFun_mono) RS subsetD) 1);
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   183
by (fast_tac (claset() addDs
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   184
		[FiniteFun_domain_Fin, Fin.dom_subset RS subsetD]) 1);
8183
344888de76c4 expandshort
paulson
parents: 8181
diff changeset
   185
by (rtac fun_FiniteFunI 1);
344888de76c4 expandshort
paulson
parents: 8181
diff changeset
   186
by (etac FiniteFun_domain_Fin 1);
8181
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   187
by (res_inst_tac [("B" , "range(f)")] fun_weaken_type 1);
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   188
by (ALLGOALS
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   189
    (blast_tac (claset() addDs
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   190
		[FiniteFun_is_fun, range_of_fun, range_type,
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   191
		 apply_equality])));
ee74d3843214 new theorems by Sidi O. Ehmety
paulson
parents: 6070
diff changeset
   192
qed "FiniteFun_Collect_iff";