src/ZF/AC/recfunAC16.ML
author wenzelm
Mon, 16 Dec 1996 10:02:17 +0100
changeset 2404 edcc26b1461d
parent 1461 6bcb44e4d6e5
child 2469 b50b8c0eec01
permissions -rw-r--r--
added symbol_input.ML;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
     1
(*  Title:      ZF/AC/recfunAC16.ML
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
     3
    Author:     Krzysztof Grabczewski
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
     4
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
     5
Properties of the recursive definition used in the proof of WO2 ==> AC16
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
     6
*)
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
     7
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
     8
open recfunAC16;
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
     9
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    10
(* ********************************************************************** *)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    11
(* Basic properties of recfunAC16                                         *)
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    12
(* ********************************************************************** *)
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    13
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    14
goalw thy [recfunAC16_def] "recfunAC16(f,fa,0,a) = 0";
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1196
diff changeset
    15
by (rtac transrec2_0 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    16
val recfunAC16_0 = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    17
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    18
goalw thy [recfunAC16_def] "recfunAC16(f,fa,succ(i),a) =  \
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    19
\       if (EX y:recfunAC16(f,fa,i,a). fa ` i <= y, recfunAC16(f,fa,i,a), \
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    20
\       recfunAC16(f,fa,i,a) Un {f ` (LEAST j. fa ` i <= f ` j &  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    21
\       (ALL b<a. (fa`b <= f`j  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    22
\       --> (ALL t:recfunAC16(f,fa,i,a). ~ fa`b <= t))))})";
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1196
diff changeset
    23
by (rtac transrec2_succ 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    24
val recfunAC16_succ = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    25
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    26
goalw thy [recfunAC16_def] "!!i. Limit(i)  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    27
\       ==> recfunAC16(f,fa,i,a) = (UN j<i. recfunAC16(f,fa,j,a))";
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1196
diff changeset
    28
by (etac transrec2_Limit 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    29
val recfunAC16_Limit = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    30
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    31
(* ********************************************************************** *)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    32
(* Monotonicity of transrec2                                              *)
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    33
(* ********************************************************************** *)
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    34
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    35
val [prem1, prem2] = goal thy 
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    36
    "[| !!g r. r <= B(g,r); Ord(i) |]  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    37
\       ==> j<i --> transrec2(j, 0, B) <= transrec2(i, 0, B)";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    38
by (resolve_tac [prem2 RS trans_induct] 1);
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1196
diff changeset
    39
by (rtac Ord_cases 1 THEN (REPEAT (assume_tac 1)));
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    40
by (fast_tac lt_cs 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    41
by (asm_simp_tac (AC_ss addsimps [transrec2_succ]) 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    42
by (fast_tac (FOL_cs addSIs [succI1, prem1]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    43
        addSEs [ballE, leE, prem1 RSN (2, subset_trans)]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    44
by (fast_tac (AC_cs addIs [OUN_I, ltI]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    45
        addSEs [Limit_has_succ RS ltE, succI1 RSN (2, Ord_in_Ord) RS le_refl,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    46
                transrec2_Limit RS ssubst]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    47
val transrec2_mono_lemma = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    48
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    49
val [prem1, prem2] = goal thy "[| !!g r. r <= B(g,r); j le i |]  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    50
\       ==> transrec2(j, 0, B) <= transrec2(i, 0, B)";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    51
by (resolve_tac [prem2 RS leE] 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    52
by (resolve_tac [transrec2_mono_lemma RS impE] 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    53
by (TRYALL (fast_tac (AC_cs addSIs [prem1, prem2, lt_Ord2])));
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    54
val transrec2_mono = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    55
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    56
(* ********************************************************************** *)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    57
(* Monotonicity of recfunAC16                                             *)
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    58
(* ********************************************************************** *)
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    59
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    60
goalw thy [recfunAC16_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    61
        "!!i. i le j ==> recfunAC16(f, g, i, a) <= recfunAC16(f, g, j, a)";
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1196
diff changeset
    62
by (rtac transrec2_mono 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    63
by (REPEAT (fast_tac (AC_cs addIs [expand_if RS iffD2]) 1));
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    64
val recfunAC16_mono = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    65