src/CCL/wfd.ML
author berghofe
Fri, 28 Apr 2006 15:59:31 +0200
changeset 19497 630073ef9212
parent 17456 bcf7544875b2
permissions -rw-r--r--
Added new targets for nominal datatype package.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
     1
(*  Title:      CCL/Wfd.ML
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
     3
*)
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
     4
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
     5
(***********)
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
     6
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
     7
val [major,prem] = goalw (the_context ()) [Wfd_def]
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
     8
    "[| Wfd(R);       \
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
     9
\       !!x.[| ALL y. <y,x>: R --> P(y) |] ==> P(x) |]  ==>  \
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    10
\    P(a)";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    11
by (rtac (major RS spec RS mp RS spec RS CollectD) 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    12
by (fast_tac (set_cs addSIs [prem RS CollectI]) 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    13
qed "wfd_induct";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    14
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    15
val [p1,p2,p3] = goal (the_context ())
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    16
    "[| !!x y.<x,y> : R ==> Q(x); \
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    17
\       ALL x. (ALL y. <y,x> : R --> y : P) --> x : P; \
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    18
\       !!x. Q(x) ==> x:P |] ==> a:P";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    19
by (rtac (p2 RS  spec  RS mp) 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    20
by (fast_tac (set_cs addSIs [p1 RS p3]) 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    21
qed "wfd_strengthen_lemma";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    22
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    23
fun wfd_strengthen_tac s i = res_inst_tac [("Q",s)] wfd_strengthen_lemma i THEN
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    24
                             assume_tac (i+1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    25
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    26
val wfd::prems = goal (the_context ()) "[| Wfd(r);  <a,x>:r;  <x,a>:r |] ==> P";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    27
by (subgoal_tac "ALL x. <a,x>:r --> <x,a>:r --> P" 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    28
by (fast_tac (FOL_cs addIs prems) 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    29
by (rtac (wfd RS  wfd_induct) 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    30
by (ALLGOALS (fast_tac (ccl_cs addSIs prems)));
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    31
qed "wf_anti_sym";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    32
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    33
val prems = goal (the_context ()) "[| Wfd(r);  <a,a>: r |] ==> P";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    34
by (rtac wf_anti_sym 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    35
by (REPEAT (resolve_tac prems 1));
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    36
qed "wf_anti_refl";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    37
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    38
(*** Irreflexive transitive closure ***)
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    39
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    40
val [prem] = goal (the_context ()) "Wfd(R) ==> Wfd(R^+)";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    41
by (rewtac Wfd_def);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    42
by (REPEAT (ares_tac [allI,ballI,impI] 1));
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    43
(*must retain the universal formula for later use!*)
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    44
by (rtac allE 1 THEN assume_tac 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    45
by (etac mp 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    46
by (rtac (prem RS wfd_induct) 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    47
by (rtac (impI RS allI) 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    48
by (etac tranclE 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    49
by (fast_tac ccl_cs 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    50
by (etac (spec RS mp RS spec RS mp) 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    51
by (REPEAT (atac 1));
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    52
qed "trancl_wf";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    53
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    54
(*** Lexicographic Ordering ***)
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    55
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    56
Goalw [lex_def]
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    57
 "p : ra**rb <-> (EX a a' b b'. p = <<a,b>,<a',b'>> & (<a,a'> : ra | a=a' & <b,b'> : rb))";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    58
by (fast_tac ccl_cs 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    59
qed "lexXH";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    60
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    61
val prems = goal (the_context ())
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    62
 "<a,a'> : ra ==> <<a,b>,<a',b'>> : ra**rb";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    63
by (fast_tac (ccl_cs addSIs (prems @ [lexXH RS iffD2])) 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    64
qed "lexI1";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    65
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    66
val prems = goal (the_context ())
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    67
 "<b,b'> : rb ==> <<a,b>,<a,b'>> : ra**rb";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    68
by (fast_tac (ccl_cs addSIs (prems @ [lexXH RS iffD2])) 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    69
qed "lexI2";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    70
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    71
val major::prems = goal (the_context ())
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    72
 "[| p : ra**rb;  \
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    73
\    !!a a' b b'.[| <a,a'> : ra; p=<<a,b>,<a',b'>> |] ==> R;  \
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    74
\    !!a b b'.[| <b,b'> : rb;  p = <<a,b>,<a,b'>> |] ==> R  |] ==> \
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    75
\ R";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    76
by (rtac (major RS (lexXH RS iffD1) RS exE) 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    77
by (REPEAT_SOME (eresolve_tac ([exE,conjE,disjE]@prems)));
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    78
by (ALLGOALS (fast_tac ccl_cs));
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    79
qed "lexE";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    80
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    81
val [major,minor] = goal (the_context ())
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    82
 "[| p : r**s;  !!a a' b b'. p = <<a,b>,<a',b'>> ==> P |] ==>P";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    83
by (rtac (major RS lexE) 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    84
by (ALLGOALS (fast_tac (set_cs addSEs [minor])));
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    85
qed "lex_pair";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    86
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    87
val [wfa,wfb] = goal (the_context ())
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    88
 "[| Wfd(R); Wfd(S) |] ==> Wfd(R**S)";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    89
by (rewtac Wfd_def);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    90
by (safe_tac ccl_cs);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    91
by (wfd_strengthen_tac "%x. EX a b. x=<a,b>" 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    92
by (fast_tac (term_cs addSEs [lex_pair]) 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    93
by (subgoal_tac "ALL a b.<a,b>:P" 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    94
by (fast_tac ccl_cs 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    95
by (rtac (wfa RS wfd_induct RS allI) 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    96
by (rtac (wfb RS wfd_induct RS allI) 1);back();
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    97
by (fast_tac (type_cs addSEs [lexE]) 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    98
qed "lex_wf";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
    99
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   100
(*** Mapping ***)
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   101
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   102
Goalw [wmap_def]
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   103
 "p : wmap(f,r) <-> (EX x y. p=<x,y>  &  <f(x),f(y)> : r)";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   104
by (fast_tac ccl_cs 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   105
qed "wmapXH";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   106
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   107
val prems = goal (the_context ())
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   108
 "<f(a),f(b)> : r ==> <a,b> : wmap(f,r)";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   109
by (fast_tac (ccl_cs addSIs (prems @ [wmapXH RS iffD2])) 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   110
qed "wmapI";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   111
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   112
val major::prems = goal (the_context ())
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   113
 "[| p : wmap(f,r);  !!a b.[| <f(a),f(b)> : r;  p=<a,b> |] ==> R |] ==> R";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   114
by (rtac (major RS (wmapXH RS iffD1) RS exE) 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   115
by (REPEAT_SOME (eresolve_tac ([exE,conjE,disjE]@prems)));
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   116
by (ALLGOALS (fast_tac ccl_cs));
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   117
qed "wmapE";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   118
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   119
val [wf] = goal (the_context ())
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   120
 "Wfd(r) ==> Wfd(wmap(f,r))";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   121
by (rewtac Wfd_def);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   122
by (safe_tac ccl_cs);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   123
by (subgoal_tac "ALL b. ALL a. f(a)=b-->a:P" 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   124
by (fast_tac ccl_cs 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   125
by (rtac (wf RS wfd_induct RS allI) 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   126
by (safe_tac ccl_cs);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   127
by (etac (spec RS mp) 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   128
by (safe_tac (ccl_cs addSEs [wmapE]));
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   129
by (etac (spec RS mp RS spec RS mp) 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   130
by (assume_tac 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   131
by (rtac refl 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   132
qed "wmap_wf";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   133
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   134
(* Projections *)
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   135
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   136
val prems = goal (the_context ()) "<xa,ya> : r ==> <<xa,xb>,<ya,yb>> : wmap(fst,r)";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   137
by (rtac wmapI 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   138
by (simp_tac (term_ss addsimps prems) 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   139
qed "wfstI";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   140
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   141
val prems = goal (the_context ()) "<xb,yb> : r ==> <<xa,xb>,<ya,yb>> : wmap(snd,r)";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   142
by (rtac wmapI 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   143
by (simp_tac (term_ss addsimps prems) 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   144
qed "wsndI";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   145
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   146
val prems = goal (the_context ()) "<xc,yc> : r ==> <<xa,<xb,xc>>,<ya,<yb,yc>>> : wmap(thd,r)";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   147
by (rtac wmapI 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   148
by (simp_tac (term_ss addsimps prems) 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   149
qed "wthdI";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   150
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   151
(*** Ground well-founded relations ***)
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   152
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   153
val prems = goalw (the_context ()) [wf_def]
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   154
    "[| Wfd(r);  a : r |] ==> a : wf(r)";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   155
by (fast_tac (set_cs addSIs prems) 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   156
qed "wfI";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   157
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   158
val prems = goalw (the_context ()) [Wfd_def] "Wfd({})";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   159
by (fast_tac (set_cs addEs [EmptyXH RS iffD1 RS FalseE]) 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   160
qed "Empty_wf";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   161
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   162
val prems = goalw (the_context ()) [wf_def] "Wfd(wf(R))";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   163
by (res_inst_tac [("Q","Wfd(R)")] (excluded_middle RS disjE) 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   164
by (ALLGOALS (asm_simp_tac CCL_ss));
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   165
by (rtac Empty_wf 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   166
qed "wf_wf";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   167
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   168
Goalw [NatPR_def]  "p : NatPR <-> (EX x:Nat. p=<x,succ(x)>)";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   169
by (fast_tac set_cs 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   170
qed "NatPRXH";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   171
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   172
Goalw [ListPR_def]  "p : ListPR(A) <-> (EX h:A. EX t:List(A).p=<t,h$t>)";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   173
by (fast_tac set_cs 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   174
qed "ListPRXH";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   175
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   176
val NatPRI = refl RS (bexI RS (NatPRXH RS iffD2));
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   177
val ListPRI = refl RS (bexI RS (bexI RS (ListPRXH RS iffD2)));
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   178
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   179
Goalw [Wfd_def]  "Wfd(NatPR)";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   180
by (safe_tac set_cs);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   181
by (wfd_strengthen_tac "%x. x:Nat" 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   182
by (fast_tac (type_cs addSEs [XH_to_E NatPRXH]) 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   183
by (etac Nat_ind 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   184
by (ALLGOALS (fast_tac (type_cs addEs [XH_to_E NatPRXH])));
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   185
qed "NatPR_wf";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   186
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   187
Goalw [Wfd_def]  "Wfd(ListPR(A))";
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   188
by (safe_tac set_cs);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   189
by (wfd_strengthen_tac "%x. x:List(A)" 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   190
by (fast_tac (type_cs addSEs [XH_to_E ListPRXH]) 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   191
by (etac List_ind 1);
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   192
by (ALLGOALS (fast_tac (type_cs addEs [XH_to_E ListPRXH])));
bcf7544875b2 converted to Isar theory format;
wenzelm
parents:
diff changeset
   193
qed "ListPR_wf";