src/ZF/ex/misc.ML
author paulson
Mon, 07 Sep 1998 10:45:45 +0200
changeset 5431 d50c2783f941
parent 5325 f7a5e06adea1
child 8266 4bc79ed1233b
permissions -rw-r--r--
new example
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
     1
(*  Title:      ZF/ex/misc
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1993  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
Miscellaneous examples for Zermelo-Fraenkel Set Theory 
1646
f48fafc18a88 updated comments
paulson
parents: 1461
diff changeset
     7
Composition of homomorphisms, Pastre's examples, ...
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     9
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
writeln"ZF/ex/misc";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    11
5431
d50c2783f941 new example
paulson
parents: 5325
diff changeset
    12
(*trivial example of term synthesis: apparently hard for some provers!*)
d50c2783f941 new example
paulson
parents: 5325
diff changeset
    13
Goal "a ~= b ==> a:?X & b ~: ?X";
d50c2783f941 new example
paulson
parents: 5325
diff changeset
    14
by (Blast_tac 1);
d50c2783f941 new example
paulson
parents: 5325
diff changeset
    15
result();
d50c2783f941 new example
paulson
parents: 5325
diff changeset
    16
3000
7ecb8b6a3d7f Moved blast_tac demo from ZF/func.ML to ZF/ex/misc.ML
paulson
parents: 2496
diff changeset
    17
(*Nice Blast_tac benchmark.  Proved in 0.3s; old tactics can't manage it!*)
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5068
diff changeset
    18
Goal "ALL x:S. ALL y:S. x<=y ==> EX z. S <= {z}";
3000
7ecb8b6a3d7f Moved blast_tac demo from ZF/func.ML to ZF/ex/misc.ML
paulson
parents: 2496
diff changeset
    19
by (Blast_tac 1);
7ecb8b6a3d7f Moved blast_tac demo from ZF/func.ML to ZF/ex/misc.ML
paulson
parents: 2496
diff changeset
    20
result();
7ecb8b6a3d7f Moved blast_tac demo from ZF/func.ML to ZF/ex/misc.ML
paulson
parents: 2496
diff changeset
    21
4322
5f5df208b0e0 New example
paulson
parents: 4152
diff changeset
    22
(*variant of the benchmark above*)
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5068
diff changeset
    23
Goal "ALL x:S. Union(S) <= x ==> EX z. S <= {z}";
4322
5f5df208b0e0 New example
paulson
parents: 4152
diff changeset
    24
by (Blast_tac 1);
5f5df208b0e0 New example
paulson
parents: 4152
diff changeset
    25
result();
5f5df208b0e0 New example
paulson
parents: 4152
diff changeset
    26
4109
b131edcfeac3 isatool fixclasimp;
wenzelm
parents: 4091
diff changeset
    27
context Perm.thy;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    28
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    29
(*Example 12 (credited to Peter Andrews) from
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    30
 W. Bledsoe.  A Maximal Method for Set Variables in Automatic Theorem-proving.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    31
 In: J. Hayes and D. Michie and L. Mikulich, eds.  Machine Intelligence 9.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    32
 Ellis Horwood, 53-100 (1979). *)
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4595
diff changeset
    33
Goal "(ALL F. {x}: F --> {y}:F) --> (ALL A. x:A --> y:A)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1782
diff changeset
    34
by (Best_tac 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    35
result();
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    36
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    37
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    38
(*** Composition of homomorphisms is a homomorphism ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    39
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    40
(*Given as a challenge problem in
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    41
  R. Boyer et al.,
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    42
  Set Theory in First-Order Logic: Clauses for G\"odel's Axioms,
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    43
  JAR 2 (1986), 287-327 
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    44
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    45
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
    46
(*collecting the relevant lemmas*)
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1782
diff changeset
    47
Addsimps [comp_fun, SigmaI, apply_funtype];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    48
736
a8d1758bb113 moved Cantors theorem to ZF/ZF.ML and ZF/Perm.ML
lcp
parents: 695
diff changeset
    49
(*This version uses a super application of simp_tac.  Needs setloop to help
a8d1758bb113 moved Cantors theorem to ZF/ZF.ML and ZF/Perm.ML
lcp
parents: 695
diff changeset
    50
  proving conditions of rewrites such as comp_fun_apply;
a8d1758bb113 moved Cantors theorem to ZF/ZF.ML and ZF/Perm.ML
lcp
parents: 695
diff changeset
    51
  rewriting does not instantiate Vars*)
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
    52
goal Perm.thy
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    53
    "(ALL A f B g. hom(A,f,B,g) = \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    54
\          {H: A->B. f:A*A->A & g:B*B->B & \
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
    55
\                    (ALL x:A. ALL y:A. H`(f`<x,y>) = g`<H`x,H`y>)}) --> \
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    56
\    J : hom(A,f,B,g) & K : hom(B,g,C,h) -->  \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    57
\    (K O J) : hom(A,f,C,h)";
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4109
diff changeset
    58
by (asm_simp_tac (simpset() setloop (K Safe_tac)) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    59
val comp_homs = result();
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    60
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
    61
(*This version uses meta-level rewriting, safe_tac and asm_simp_tac*)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    62
val [hom_def] = goal Perm.thy
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    63
    "(!! A f B g. hom(A,f,B,g) == \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    64
\          {H: A->B. f:A*A->A & g:B*B->B & \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    65
\                    (ALL x:A. ALL y:A. H`(f`<x,y>) = g`<H`x,H`y>)}) ==> \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    66
\    J : hom(A,f,B,g) & K : hom(B,g,C,h) -->  \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    67
\    (K O J) : hom(A,f,C,h)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    68
by (rewtac hom_def);
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4109
diff changeset
    69
by Safe_tac;
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1782
diff changeset
    70
by (Asm_simp_tac 1);
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1782
diff changeset
    71
by (Asm_simp_tac 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 736
diff changeset
    72
qed "comp_homs";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    73
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    74
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    75
(** A characterization of functions, suggested by Tobias Nipkow **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    76
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4595
diff changeset
    77
Goalw [Pi_def, function_def]
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    78
    "r: domain(r)->B  <->  r <= domain(r)*B & (ALL X. r `` (r -`` X) <= X)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1782
diff changeset
    79
by (Best_tac 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    80
result();
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    81
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    82
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    83
(**** From D Pastre.  Automatic theorem proving in set theory. 
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    84
         Artificial Intelligence, 10:1--27, 1978.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    85
             These examples require forward reasoning! ****)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    86
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    87
(*reduce the clauses to units by type checking -- beware of nontermination*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    88
fun forw_typechk tyrls [] = []
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    89
  | forw_typechk tyrls clauses =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    90
    let val (units, others) = partition (has_fewer_prems 1) clauses
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    91
    in  gen_union eq_thm (units, forw_typechk tyrls (tyrls RL others))
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    92
    end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    93
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    94
(*A crude form of forward reasoning*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    95
fun forw_iterate tyrls rls facts 0 = facts
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    96
  | forw_iterate tyrls rls facts n =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    97
      let val facts' = 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
    98
          gen_union eq_thm (forw_typechk (tyrls@facts) (facts RL rls), facts);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    99
      in  forw_iterate tyrls rls facts' (n-1)  end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   100
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   101
val pastre_rls =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   102
    [comp_mem_injD1, comp_mem_surjD1, comp_mem_injD2, comp_mem_surjD2];
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   103
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   104
fun pastre_facts (fact1::fact2::fact3::prems) = 
434
89d45187f04d Various updates and tidying
lcp
parents: 38
diff changeset
   105
    forw_iterate (prems @ [comp_surj, comp_inj, comp_fun])
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   106
               pastre_rls [fact1,fact2,fact3] 4;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   107
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   108
val prems = goalw Perm.thy [bij_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   109
    "[| (h O g O f): inj(A,A);          \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   110
\       (f O h O g): surj(B,B);         \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   111
\       (g O f O h): surj(C,C);         \
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   112
\       f: A->B;  g: B->C;  h: C->A |] ==> h: bij(C,A)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   113
by (REPEAT (resolve_tac (IntI :: pastre_facts prems) 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 736
diff changeset
   114
qed "pastre1";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   115
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   116
val prems = goalw Perm.thy [bij_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   117
    "[| (h O g O f): surj(A,A);         \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   118
\       (f O h O g): inj(B,B);          \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   119
\       (g O f O h): surj(C,C);         \
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   120
\       f: A->B;  g: B->C;  h: C->A |] ==> h: bij(C,A)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   121
by (REPEAT (resolve_tac (IntI :: pastre_facts prems) 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 736
diff changeset
   122
qed "pastre2";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   123
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   124
val prems = goalw Perm.thy [bij_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   125
    "[| (h O g O f): surj(A,A);         \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   126
\       (f O h O g): surj(B,B);         \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   127
\       (g O f O h): inj(C,C);          \
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   128
\       f: A->B;  g: B->C;  h: C->A |] ==> h: bij(C,A)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   129
by (REPEAT (resolve_tac (IntI :: pastre_facts prems) 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 736
diff changeset
   130
qed "pastre3";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   131
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   132
val prems = goalw Perm.thy [bij_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   133
    "[| (h O g O f): surj(A,A);         \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   134
\       (f O h O g): inj(B,B);          \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   135
\       (g O f O h): inj(C,C);          \
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   136
\       f: A->B;  g: B->C;  h: C->A |] ==> h: bij(C,A)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   137
by (REPEAT (resolve_tac (IntI :: pastre_facts prems) 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 736
diff changeset
   138
qed "pastre4";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   139
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   140
val prems = goalw Perm.thy [bij_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   141
    "[| (h O g O f): inj(A,A);          \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   142
\       (f O h O g): surj(B,B);         \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   143
\       (g O f O h): inj(C,C);          \
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   144
\       f: A->B;  g: B->C;  h: C->A |] ==> h: bij(C,A)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   145
by (REPEAT (resolve_tac (IntI :: pastre_facts prems) 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 736
diff changeset
   146
qed "pastre5";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   147
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   148
val prems = goalw Perm.thy [bij_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   149
    "[| (h O g O f): inj(A,A);          \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   150
\       (f O h O g): inj(B,B);          \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1110
diff changeset
   151
\       (g O f O h): surj(C,C);         \
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   152
\       f: A->B;  g: B->C;  h: C->A |] ==> h: bij(C,A)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   153
by (REPEAT (resolve_tac (IntI :: pastre_facts prems) 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 736
diff changeset
   154
qed "pastre6";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   155
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
   156
(** Yet another example... **)
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
   157
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1782
diff changeset
   158
goal Perm.thy
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
   159
    "(lam Z:Pow(A+B). <{x:A. Inl(x):Z}, {y:B. Inr(y):Z}>) \
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
   160
\    : bij(Pow(A+B), Pow(A)*Pow(B))";
1110
756aa2e81f6e Changed some definitions and proofs to use pattern-matching.
lcp
parents: 782
diff changeset
   161
by (res_inst_tac [("d", "%<X,Y>.{Inl(x).x:X} Un {Inr(y).y:Y}")] 
695
a1586fa1b755 ZF/ex/Ramsey,Rmap,misc.ML: modified for new definition of Pi(A,B)
lcp
parents: 434
diff changeset
   162
    lam_bijective 1);
4477
b3e5857d8d99 New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents: 4322
diff changeset
   163
(*Auto_tac no longer proves it*)
b3e5857d8d99 New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents: 4322
diff changeset
   164
by (REPEAT (fast_tac (claset() addss (simpset())) 1));
4595
fa8cee619732 New example, Pow_Sigma_bij
paulson
parents: 4477
diff changeset
   165
qed "Pow_sum_bij";
fa8cee619732 New example, Pow_Sigma_bij
paulson
parents: 4477
diff changeset
   166
fa8cee619732 New example, Pow_Sigma_bij
paulson
parents: 4477
diff changeset
   167
(*As a special case, we have  bij(Pow(A*B), A -> Pow B)  *)
fa8cee619732 New example, Pow_Sigma_bij
paulson
parents: 4477
diff changeset
   168
goal Perm.thy
fa8cee619732 New example, Pow_Sigma_bij
paulson
parents: 4477
diff changeset
   169
    "(lam r:Pow(Sigma(A,B)). lam x:A. r``{x}) \
fa8cee619732 New example, Pow_Sigma_bij
paulson
parents: 4477
diff changeset
   170
\    : bij(Pow(Sigma(A,B)), PROD x:A. Pow(B(x)))";
fa8cee619732 New example, Pow_Sigma_bij
paulson
parents: 4477
diff changeset
   171
by (res_inst_tac [("d", "%f. UN x:A. UN y:f`x. {<x,y>}")] lam_bijective 1);
fa8cee619732 New example, Pow_Sigma_bij
paulson
parents: 4477
diff changeset
   172
by (blast_tac (claset() addDs [apply_type]) 2);
fa8cee619732 New example, Pow_Sigma_bij
paulson
parents: 4477
diff changeset
   173
by (blast_tac (claset() addIs [lam_type]) 1);
fa8cee619732 New example, Pow_Sigma_bij
paulson
parents: 4477
diff changeset
   174
by (ALLGOALS Asm_simp_tac);
fa8cee619732 New example, Pow_Sigma_bij
paulson
parents: 4477
diff changeset
   175
by (Fast_tac 1);
fa8cee619732 New example, Pow_Sigma_bij
paulson
parents: 4477
diff changeset
   176
by (rtac fun_extension 1);
fa8cee619732 New example, Pow_Sigma_bij
paulson
parents: 4477
diff changeset
   177
by (assume_tac 2);
fa8cee619732 New example, Pow_Sigma_bij
paulson
parents: 4477
diff changeset
   178
by (rtac (singletonI RS lam_type) 1);
fa8cee619732 New example, Pow_Sigma_bij
paulson
parents: 4477
diff changeset
   179
by (Asm_simp_tac 1);
fa8cee619732 New example, Pow_Sigma_bij
paulson
parents: 4477
diff changeset
   180
by (Blast_tac 1);
fa8cee619732 New example, Pow_Sigma_bij
paulson
parents: 4477
diff changeset
   181
qed "Pow_Sigma_bij";
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
   182
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   183
writeln"Reached end of file.";