src/ZF/WF.ML
author paulson
Fri, 18 Feb 2000 15:35:29 +0100
changeset 8255 38f96394c099
parent 7570 a9391550eea1
child 9173 422968aeed49
permissions -rw-r--r--
new distributive laws
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
     1
(*  Title:      ZF/wf.ML
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
     3
    Author:     Tobias Nipkow and Lawrence C Paulson
4515
44af72721564 Now calls Blast_tac more often
paulson
parents: 4091
diff changeset
     4
    Copyright   1998  University of Cambridge
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
4515
44af72721564 Now calls Blast_tac more often
paulson
parents: 4091
diff changeset
     6
Well-founded Recursion
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
Derived first for transitive relations, and finally for arbitrary WF relations
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     9
via wf_trancl and trans_trancl.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    11
It is difficult to derive this general case directly, using r^+ instead of
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    12
r.  In is_recfun, the two occurrences of the relation must have the same
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    13
form.  Inserting r^+ in the_recfun or wftrec yields a recursion rule with
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    14
r^+ -`` {a} instead of r-``{a}.  This recursion rule is stronger in
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    15
principle, but harder to use, especially to prove wfrec_eclose_eq in
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    16
epsilon.ML.  Expanding out the definition of wftrec in wfrec would yield
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    17
a mess.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    18
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    19
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    20
open WF;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    21
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    22
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    23
(*** Well-founded relations ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    24
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
    25
(** Equivalences between wf and wf_on **)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
    26
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
    27
Goalw [wf_def, wf_on_def] "wf(r) ==> wf[A](r)";
4515
44af72721564 Now calls Blast_tac more often
paulson
parents: 4091
diff changeset
    28
by (Clarify_tac 1);  (*essential for Blast_tac's efficiency*)
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2637
diff changeset
    29
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 494
diff changeset
    30
qed "wf_imp_wf_on";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
    31
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
    32
Goalw [wf_def, wf_on_def] "wf[field(r)](r) ==> wf(r)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
    33
by (Fast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 494
diff changeset
    34
qed "wf_on_field_imp_wf";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
    35
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4515
diff changeset
    36
Goal "wf(r) <-> wf[field(r)](r)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3016
diff changeset
    37
by (blast_tac (claset() addIs [wf_imp_wf_on, wf_on_field_imp_wf]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 494
diff changeset
    38
qed "wf_iff_wf_on_field";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    39
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
    40
Goalw [wf_on_def, wf_def] "[| wf[A](r);  B<=A |] ==> wf[B](r)";
5265
9d1d4c43c76d Disjointness reasoning by AddEs [equals0E, sym RS equals0E]
paulson
parents: 5147
diff changeset
    41
by (Fast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 494
diff changeset
    42
qed "wf_on_subset_A";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
    43
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
    44
Goalw [wf_on_def, wf_def] "[| wf[A](r);  s<=r |] ==> wf[A](s)";
5265
9d1d4c43c76d Disjointness reasoning by AddEs [equals0E, sym RS equals0E]
paulson
parents: 5147
diff changeset
    45
by (Fast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 494
diff changeset
    46
qed "wf_on_subset_r";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
    47
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
    48
(** Introduction rules for wf_on **)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
    49
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
    50
(*If every non-empty subset of A has an r-minimal element then wf[A](r).*)
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5265
diff changeset
    51
val [prem] = Goalw [wf_on_def, wf_def]
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
    52
    "[| !!Z u. [| Z<=A;  u:Z;  ALL x:Z. EX y:Z. <y,x>:r |] ==> False |] \
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
    53
\    ==>  wf[A](r)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    54
by (rtac (equals0I RS disjCI RS allI) 1);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
    55
by (res_inst_tac [ ("Z", "Z") ] prem 1);
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2637
diff changeset
    56
by (ALLGOALS Blast_tac);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 494
diff changeset
    57
qed "wf_onI";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    58
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
    59
(*If r allows well-founded induction over A then wf[A](r)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
    60
  Premise is equivalent to 
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
    61
  !!B. ALL x:A. (ALL y. <y,x>: r --> y:B) --> x:B ==> A<=B  *)
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5265
diff changeset
    62
val [prem] = Goal
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
    63
    "[| !!y B. [| ALL x:A. (ALL y:A. <y,x>:r --> y:B) --> x:B;   y:A  \
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
    64
\              |] ==> y:B |] \
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
    65
\    ==>  wf[A](r)";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
    66
by (rtac wf_onI 1);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
    67
by (res_inst_tac [ ("c", "u") ] (prem RS DiffE) 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
    68
by (contr_tac 3);
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2637
diff changeset
    69
by (Blast_tac 2);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
    70
by (Fast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 494
diff changeset
    71
qed "wf_onI2";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    72
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    73
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    74
(** Well-founded Induction **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    75
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    76
(*Consider the least z in domain(r) Un {a} such that P(z) does not hold...*)
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5265
diff changeset
    77
val [major,minor] = Goalw [wf_def]
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    78
    "[| wf(r);          \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    79
\       !!x.[| ALL y. <y,x>: r --> P(y) |] ==> P(x) \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    80
\    |]  ==>  P(a)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    81
by (res_inst_tac [ ("x", "{z:domain(r) Un {a}. ~P(z)}") ]  (major RS allE) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    82
by (etac disjE 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3016
diff changeset
    83
by (blast_tac (claset() addEs [equalityE]) 1);
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3016
diff changeset
    84
by (asm_full_simp_tac (simpset() addsimps [domainI]) 1);
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3016
diff changeset
    85
by (blast_tac (claset() addSDs [minor]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 494
diff changeset
    86
qed "wf_induct";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    87
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    88
(*Perform induction on i, then prove the wf(r) subgoal using prems. *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    89
fun wf_ind_tac a prems i = 
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    90
    EVERY [res_inst_tac [("a",a)] wf_induct i,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
    91
           rename_last_tac a ["1"] (i+1),
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
    92
           ares_tac prems i];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    93
485
5e00a676a211 Axiom of choice, cardinality results, etc.
lcp
parents: 443
diff changeset
    94
(*The form of this rule is designed to match wfI*)
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5265
diff changeset
    95
val wfr::amem::prems = Goal
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    96
    "[| wf(r);  a:A;  field(r)<=A;  \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    97
\       !!x.[| x: A;  ALL y. <y,x>: r --> P(y) |] ==> P(x) \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    98
\    |]  ==>  P(a)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    99
by (rtac (amem RS rev_mp) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   100
by (wf_ind_tac "a" [wfr] 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   101
by (rtac impI 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   102
by (eresolve_tac prems 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3016
diff changeset
   103
by (blast_tac (claset() addIs (prems RL [subsetD])) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 494
diff changeset
   104
qed "wf_induct2";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   105
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5265
diff changeset
   106
Goal "!!r A. field(r Int A*A) <= A";
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2637
diff changeset
   107
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 494
diff changeset
   108
qed "field_Int_square";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   109
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5265
diff changeset
   110
val wfr::amem::prems = Goalw [wf_on_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
   111
    "[| wf[A](r);  a:A;                                         \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
   112
\       !!x.[| x: A;  ALL y:A. <y,x>: r --> P(y) |] ==> P(x)    \
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   113
\    |]  ==>  P(a)";
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   114
by (rtac ([wfr, amem, field_Int_square] MRS wf_induct2) 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   115
by (REPEAT (ares_tac prems 1));
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2637
diff changeset
   116
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 494
diff changeset
   117
qed "wf_on_induct";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   118
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   119
fun wf_on_ind_tac a prems i = 
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   120
    EVERY [res_inst_tac [("a",a)] wf_on_induct i,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
   121
           rename_last_tac a ["1"] (i+2),
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
   122
           REPEAT (ares_tac prems i)];
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   123
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   124
(*If r allows well-founded induction then wf(r)*)
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5265
diff changeset
   125
val [subs,indhyp] = Goal
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   126
    "[| field(r)<=A;  \
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   127
\       !!y B. [| ALL x:A. (ALL y:A. <y,x>:r --> y:B) --> x:B;   y:A  \
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   128
\              |] ==> y:B |] \
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   129
\    ==>  wf(r)";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   130
by (rtac ([wf_onI2, subs] MRS (wf_on_subset_A RS wf_on_field_imp_wf)) 1);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   131
by (REPEAT (ares_tac [indhyp] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 494
diff changeset
   132
qed "wfI";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   133
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   134
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   135
(*** Properties of well-founded relations ***)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   136
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   137
Goal "wf(r) ==> <a,a> ~: r";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   138
by (wf_ind_tac "a" [] 1);
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2637
diff changeset
   139
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 494
diff changeset
   140
qed "wf_not_refl";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   141
5452
b38332431a8c New theorem wf_not_sym and well-formed wf_asym
paulson
parents: 5321
diff changeset
   142
Goal "wf(r) ==> ALL x. <a,x>:r --> <x,a> ~: r";
b38332431a8c New theorem wf_not_sym and well-formed wf_asym
paulson
parents: 5321
diff changeset
   143
by (wf_ind_tac "a" [] 1);
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2637
diff changeset
   144
by (Blast_tac 1);
5452
b38332431a8c New theorem wf_not_sym and well-formed wf_asym
paulson
parents: 5321
diff changeset
   145
qed_spec_mp "wf_not_sym";
b38332431a8c New theorem wf_not_sym and well-formed wf_asym
paulson
parents: 5321
diff changeset
   146
b38332431a8c New theorem wf_not_sym and well-formed wf_asym
paulson
parents: 5321
diff changeset
   147
(* [| wf(r);  <a,x> : r;  ~P ==> <x,a> : r |] ==> P *)
b38332431a8c New theorem wf_not_sym and well-formed wf_asym
paulson
parents: 5321
diff changeset
   148
bind_thm ("wf_asym", wf_not_sym RS swap);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   149
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   150
Goal "[| wf[A](r); a: A |] ==> <a,a> ~: r";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   151
by (wf_on_ind_tac "a" [] 1);
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2637
diff changeset
   152
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 494
diff changeset
   153
qed "wf_on_not_refl";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   154
5452
b38332431a8c New theorem wf_not_sym and well-formed wf_asym
paulson
parents: 5321
diff changeset
   155
Goal "[| wf[A](r);  a:A;  b:A |] ==> <a,b>:r --> <b,a>~:r";
b38332431a8c New theorem wf_not_sym and well-formed wf_asym
paulson
parents: 5321
diff changeset
   156
by (res_inst_tac [("x","b")] bspec 1);
b38332431a8c New theorem wf_not_sym and well-formed wf_asym
paulson
parents: 5321
diff changeset
   157
by (assume_tac 2);
b38332431a8c New theorem wf_not_sym and well-formed wf_asym
paulson
parents: 5321
diff changeset
   158
by (wf_on_ind_tac "a" [] 1);
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2637
diff changeset
   159
by (Blast_tac 1);
5452
b38332431a8c New theorem wf_not_sym and well-formed wf_asym
paulson
parents: 5321
diff changeset
   160
qed_spec_mp "wf_on_not_sym";
b38332431a8c New theorem wf_not_sym and well-formed wf_asym
paulson
parents: 5321
diff changeset
   161
b38332431a8c New theorem wf_not_sym and well-formed wf_asym
paulson
parents: 5321
diff changeset
   162
(* [| wf[A](r);  <a,b> : r;  a:A;  b:A;  ~P ==> <b,a> : r |] ==> P *)
b38332431a8c New theorem wf_not_sym and well-formed wf_asym
paulson
parents: 5321
diff changeset
   163
bind_thm ("wf_on_asym", wf_on_not_sym RS swap);
b38332431a8c New theorem wf_not_sym and well-formed wf_asym
paulson
parents: 5321
diff changeset
   164
b38332431a8c New theorem wf_not_sym and well-formed wf_asym
paulson
parents: 5321
diff changeset
   165
val prems =
b38332431a8c New theorem wf_not_sym and well-formed wf_asym
paulson
parents: 5321
diff changeset
   166
Goal "[| wf[A](r);  <a,b>:r;  ~P ==> <b,a>:r;  a:A;  b:A |] ==> P";
b38332431a8c New theorem wf_not_sym and well-formed wf_asym
paulson
parents: 5321
diff changeset
   167
by (rtac ccontr 1);
b38332431a8c New theorem wf_not_sym and well-formed wf_asym
paulson
parents: 5321
diff changeset
   168
by (rtac (wf_on_not_sym RS notE) 1);
b38332431a8c New theorem wf_not_sym and well-formed wf_asym
paulson
parents: 5321
diff changeset
   169
by (DEPTH_SOLVE (ares_tac prems 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 494
diff changeset
   170
qed "wf_on_asym";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   171
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   172
(*Needed to prove well_ordI.  Could also reason that wf[A](r) means
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   173
  wf(r Int A*A);  thus wf( (r Int A*A)^+ ) and use wf_not_refl *)
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   174
Goal "[| wf[A](r); <a,b>:r; <b,c>:r; <c,a>:r; a:A; b:A; c:A |] ==> P";
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   175
by (subgoal_tac "ALL y:A. ALL z:A. <a,y>:r --> <y,z>:r --> <z,a>:r --> P" 1);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   176
by (wf_on_ind_tac "a" [] 2);
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2637
diff changeset
   177
by (Blast_tac 2);
15763781afb0 Conversion to use blast_tac
paulson
parents: 2637
diff changeset
   178
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 494
diff changeset
   179
qed "wf_on_chain3";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   180
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   181
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   182
(*retains the universal formula for later use!*)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   183
val bchain_tac = EVERY' [rtac (bspec RS mp), assume_tac, assume_tac ];
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   184
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   185
(*transitive closure of a WF relation is WF provided A is downwards closed*)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   186
val [wfr,subs] = goal WF.thy
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   187
    "[| wf[A](r);  r-``A <= A |] ==> wf[A](r^+)";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   188
by (rtac wf_onI2 1);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   189
by (bchain_tac 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   190
by (eres_inst_tac [("a","y")] (wfr RS wf_on_induct) 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   191
by (cut_facts_tac [subs] 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3016
diff changeset
   192
by (blast_tac (claset() addEs [tranclE]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 494
diff changeset
   193
qed "wf_on_trancl";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   194
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   195
Goal "wf(r) ==> wf(r^+)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3016
diff changeset
   196
by (asm_full_simp_tac (simpset() addsimps [wf_iff_wf_on_field]) 1);
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   197
by (rtac (trancl_type RS field_rel_subset RSN (2, wf_on_subset_A)) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   198
by (etac wf_on_trancl 1);
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2637
diff changeset
   199
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 494
diff changeset
   200
qed "wf_trancl";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   201
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   202
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   203
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   204
(** r-``{a} is the set of everything under a in r **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   205
6112
5e4871c5136b datatype package improvements
paulson
parents: 5452
diff changeset
   206
bind_thm ("underI", vimage_singleton_iff RS iffD2);
5e4871c5136b datatype package improvements
paulson
parents: 5452
diff changeset
   207
bind_thm ("underD", vimage_singleton_iff RS iffD1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   208
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   209
(** is_recfun **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   210
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5265
diff changeset
   211
Goalw [is_recfun_def] "is_recfun(r,a,H,f) ==> f: r-``{a} -> range(f)";
f8848433d240 got rid of some goal thy commands
paulson
parents: 5265
diff changeset
   212
by (etac ssubst 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   213
by (rtac (lamI RS rangeI RS lam_type) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   214
by (assume_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 494
diff changeset
   215
qed "is_recfun_type";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   216
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   217
val [isrec,rel] = goalw WF.thy [is_recfun_def]
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   218
    "[| is_recfun(r,a,H,f); <x,a>:r |] ==> f`x = H(x, restrict(f,r-``{x}))";
443
10884e64c241 added parentheses made necessary by new constrain precedence
clasohm
parents: 437
diff changeset
   219
by (res_inst_tac [("P", "%x.?t(x) = (?u::i)")] (isrec RS ssubst) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   220
by (rtac (rel RS underI RS beta) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 494
diff changeset
   221
qed "apply_recfun";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   222
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   223
(*eresolve_tac transD solves <a,b>:r using transitivity AT MOST ONCE
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   224
  spec RS mp  instantiates induction hypotheses*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   225
fun indhyp_tac hyps =
6112
5e4871c5136b datatype package improvements
paulson
parents: 5452
diff changeset
   226
    resolve_tac (TrueI::refl::reflexive_thm::hyps) ORELSE' 
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   227
    (cut_facts_tac hyps THEN'
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   228
       DEPTH_SOLVE_1 o (ares_tac [TrueI, ballI] ORELSE'
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
   229
                        eresolve_tac [underD, transD, spec RS mp]));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   230
6
8ce8c4d13d4d Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents: 0
diff changeset
   231
(*** NOTE! some simplifications need a different solver!! ***)
7570
a9391550eea1 Mod because of new solver interface.
nipkow
parents: 6112
diff changeset
   232
val wf_super_ss = simpset() setSolver (mk_solver "WF" indhyp_tac);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   233
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   234
Goalw [is_recfun_def]
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   235
    "[| wf(r);  trans(r);  is_recfun(r,a,H,f);  is_recfun(r,b,H,g) |] ==> \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   236
\    <x,a>:r --> <x,b>:r --> f`x=g`x";
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   237
by (wf_ind_tac "x" [] 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   238
by (REPEAT (rtac impI 1 ORELSE etac ssubst 1));
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   239
by (rewtac restrict_def);
6
8ce8c4d13d4d Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents: 0
diff changeset
   240
by (asm_simp_tac (wf_super_ss addsimps [vimage_singleton_iff]) 1);
6112
5e4871c5136b datatype package improvements
paulson
parents: 5452
diff changeset
   241
qed_spec_mp "is_recfun_equal";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   242
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   243
val prems as [wfr,transr,recf,recg,_] = goal WF.thy
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   244
    "[| wf(r);  trans(r);       \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   245
\       is_recfun(r,a,H,f);  is_recfun(r,b,H,g);  <b,a>:r |] ==> \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   246
\    restrict(f, r-``{b}) = g";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   247
by (cut_facts_tac prems 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   248
by (rtac (consI1 RS restrict_type RS fun_extension) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   249
by (etac is_recfun_type 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   250
by (ALLGOALS
6
8ce8c4d13d4d Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents: 0
diff changeset
   251
    (asm_simp_tac (wf_super_ss addsimps
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
   252
                   [ [wfr,transr,recf,recg] MRS is_recfun_equal ])));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 494
diff changeset
   253
qed "is_recfun_cut";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   254
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   255
(*** Main Existence Lemma ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   256
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5265
diff changeset
   257
Goal "[| wf(r); trans(r); is_recfun(r,a,H,f); is_recfun(r,a,H,g) |]  ==>  f=g";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   258
by (rtac fun_extension 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   259
by (REPEAT (ares_tac [is_recfun_equal] 1
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   260
     ORELSE eresolve_tac [is_recfun_type,underD] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 494
diff changeset
   261
qed "is_recfun_functional";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   262
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   263
(*If some f satisfies is_recfun(r,a,H,-) then so does the_recfun(r,a,H) *)
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5265
diff changeset
   264
Goalw [the_recfun_def]
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   265
    "[| is_recfun(r,a,H,f);  wf(r);  trans(r) |]  \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   266
\    ==> is_recfun(r, a, H, the_recfun(r,a,H))";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   267
by (rtac (ex1I RS theI) 1);
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5265
diff changeset
   268
by (REPEAT (ares_tac [is_recfun_functional] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 494
diff changeset
   269
qed "is_the_recfun";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   270
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5265
diff changeset
   271
Goal "[| wf(r);  trans(r) |] ==> is_recfun(r, a, H, the_recfun(r,a,H))";
f8848433d240 got rid of some goal thy commands
paulson
parents: 5265
diff changeset
   272
by (wf_ind_tac "a" [] 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   273
by (res_inst_tac [("f", "lam y: r-``{a1}. wftrec(r,y,H)")] is_the_recfun 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   274
by (REPEAT (assume_tac 2));
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   275
by (rewrite_goals_tac [is_recfun_def, wftrec_def]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   276
(*Applying the substitution: must keep the quantified assumption!!*)
6
8ce8c4d13d4d Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents: 0
diff changeset
   277
by (REPEAT (dtac underD 1 ORELSE resolve_tac [refl, lam_cong] 1));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   278
by (fold_tac [is_recfun_def]);
6
8ce8c4d13d4d Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents: 0
diff changeset
   279
by (rtac (consI1 RS restrict_type RSN (2,fun_extension) RS subst_context) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   280
by (rtac is_recfun_type 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   281
by (ALLGOALS
6
8ce8c4d13d4d Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents: 0
diff changeset
   282
    (asm_simp_tac
8ce8c4d13d4d Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents: 0
diff changeset
   283
     (wf_super_ss addsimps [underI RS beta, apply_recfun, is_recfun_cut])));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 494
diff changeset
   284
qed "unfold_the_recfun";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   285
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   286
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   287
(*** Unfolding wftrec ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   288
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   289
Goal "[| wf(r);  trans(r);  <b,a>:r |] ==> \
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   290
\     restrict(the_recfun(r,a,H), r-``{b}) = the_recfun(r,b,H)";
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   291
by (REPEAT (ares_tac [is_recfun_cut, unfold_the_recfun] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 494
diff changeset
   292
qed "the_recfun_cut";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   293
4515
44af72721564 Now calls Blast_tac more often
paulson
parents: 4091
diff changeset
   294
(*NOT SUITABLE FOR REWRITING: it is recursive!*)
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4515
diff changeset
   295
Goalw [wftrec_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   296
    "[| wf(r);  trans(r) |] ==> \
6
8ce8c4d13d4d Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents: 0
diff changeset
   297
\         wftrec(r,a,H) = H(a, lam x: r-``{a}. wftrec(r,x,H))";
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1461
diff changeset
   298
by (stac (rewrite_rule [is_recfun_def] unfold_the_recfun) 1);
4515
44af72721564 Now calls Blast_tac more often
paulson
parents: 4091
diff changeset
   299
by (ALLGOALS 
44af72721564 Now calls Blast_tac more often
paulson
parents: 4091
diff changeset
   300
    (asm_simp_tac
44af72721564 Now calls Blast_tac more often
paulson
parents: 4091
diff changeset
   301
     (simpset() addsimps [vimage_singleton_iff RS iff_sym, the_recfun_cut])));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 494
diff changeset
   302
qed "wftrec";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   303
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   304
(** Removal of the premise trans(r) **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   305
4515
44af72721564 Now calls Blast_tac more often
paulson
parents: 4091
diff changeset
   306
(*NOT SUITABLE FOR REWRITING: it is recursive!*)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   307
val [wfr] = goalw WF.thy [wfrec_def]
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   308
    "wf(r) ==> wfrec(r,a,H) = H(a, lam x:r-``{a}. wfrec(r,x,H))";
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1461
diff changeset
   309
by (stac (wfr RS wf_trancl RS wftrec) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   310
by (rtac trans_trancl 1);
6
8ce8c4d13d4d Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents: 0
diff changeset
   311
by (rtac (vimage_pair_mono RS restrict_lam_eq RS subst_context) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   312
by (etac r_into_trancl 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   313
by (rtac subset_refl 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 494
diff changeset
   314
qed "wfrec";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   315
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   316
(*This form avoids giant explosions in proofs.  NOTE USE OF == *)
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5265
diff changeset
   317
val rew::prems = Goal
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   318
    "[| !!x. h(x)==wfrec(r,x,H);  wf(r) |] ==> \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   319
\    h(a) = H(a, lam x: r-``{a}. h(x))";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   320
by (rewtac rew);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   321
by (REPEAT (resolve_tac (prems@[wfrec]) 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 494
diff changeset
   322
qed "def_wfrec";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   323
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5265
diff changeset
   324
val prems = Goal
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   325
    "[| wf(r);  a:A;  field(r)<=A;  \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   326
\       !!x u. [| x: A;  u: Pi(r-``{x}, B) |] ==> H(x,u) : B(x)   \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   327
\    |] ==> wfrec(r,a,H) : B(a)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   328
by (res_inst_tac [("a","a")] wf_induct2 1);
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1461
diff changeset
   329
by (stac wfrec 4);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   330
by (REPEAT (ares_tac (prems@[lam_type]) 1
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   331
     ORELSE eresolve_tac [spec RS mp, underD] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 494
diff changeset
   332
qed "wfrec_type";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   333
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   334
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4515
diff changeset
   335
Goalw [wf_on_def, wfrec_on_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   336
 "[| wf[A](r);  a: A |] ==> \
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   337
\        wfrec[A](r,a,H) = H(a, lam x: (r-``{a}) Int A. wfrec[A](r,x,H))";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   338
by (etac (wfrec RS trans) 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3016
diff changeset
   339
by (asm_simp_tac (simpset() addsimps [vimage_Int_square, cons_subset_iff]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 494
diff changeset
   340
qed "wfrec_on";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   341