src/ZF/AC/AC15_WO6.ML
author wenzelm
Tue, 07 Sep 1999 10:40:58 +0200
changeset 7499 23e090051cb8
parent 5147 825877190618
child 11317 7f9e4c389318
permissions -rw-r--r--
isatool expandshort;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1205
diff changeset
     1
(*  Title:      ZF/AC/AC15_WO6.ML
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1205
diff changeset
     3
    Author:     Krzysztof Grabczewski
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
     4
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
     5
The proof of AC1 ==> WO2
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
     6
*)
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
     7
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
     8
open AC15_WO6;
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
     9
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5068
diff changeset
    10
Goal "Ord(x) ==> (UN a<x. F(a)) = (UN a:x. F(a))";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3731
diff changeset
    11
by (fast_tac (claset() addSIs [ltI] addSDs [ltD]) 1);
3731
71366483323b result() -> qed; Step_tac -> Safe_tac
paulson
parents: 2496
diff changeset
    12
qed "OUN_eq_UN";
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    13
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    14
val [prem] = goal thy "ALL x:Pow(A)-{0}. f`x~=0 & f`x<=x & f`x lepoll m ==>  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1205
diff changeset
    15
\       (UN i<LEAST x. HH(f,A,x)={A}. HH(f,A,i)) = A";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3731
diff changeset
    16
by (simp_tac (simpset() addsimps [Ord_Least RS OUN_eq_UN]) 1);
1205
f87457b1ce5e Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    17
by (rtac equalityI 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3731
diff changeset
    18
by (fast_tac (claset() addSDs [less_Least_subset_x]) 1);
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3731
diff changeset
    19
by (fast_tac (claset() addSDs [prem RS bspec]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1205
diff changeset
    20
                addSIs [f_subsets_imp_UN_HH_eq_x RS (Diff_eq_0_iff RS iffD1)]) 1);
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    21
val lemma1 = result();
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    22
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    23
val [prem] = goal thy "ALL x:Pow(A)-{0}. f`x~=0 & f`x<=x & f`x lepoll m ==>  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1205
diff changeset
    24
\       ALL x<LEAST x. HH(f,A,x)={A}. HH(f,A,x) lepoll m";
1205
f87457b1ce5e Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    25
by (rtac oallI 1);
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    26
by (dresolve_tac [ltD RS less_Least_subset_x] 1);
7499
23e090051cb8 isatool expandshort;
wenzelm
parents: 5147
diff changeset
    27
by (ftac HH_subset_imp_eq 1);
1205
f87457b1ce5e Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    28
by (etac ssubst 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3731
diff changeset
    29
by (fast_tac (claset() addIs [prem RS ballE]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1205
diff changeset
    30
                addSDs [HH_subset_x_imp_subset_Diff_UN RS not_emptyI2]) 1);
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    31
val lemma2 = result();
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    32
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
    33
Goalw [AC15_def, WO6_def] "AC15 ==> WO6";
1205
f87457b1ce5e Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    34
by (rtac allI 1);
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    35
by (eres_inst_tac [("x","Pow(A)-{0}")] allE 1);
1205
f87457b1ce5e Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    36
by (etac impE 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    37
by (Fast_tac 1);
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    38
by (REPEAT (eresolve_tac [bexE,conjE,exE] 1));
1205
f87457b1ce5e Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    39
by (rtac bexI 1 THEN (assume_tac 2));
f87457b1ce5e Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    40
by (rtac conjI 1 THEN (assume_tac 1));
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    41
by (res_inst_tac [("x","LEAST i. HH(f,A,i)={A}")] exI 1);
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    42
by (res_inst_tac [("x","lam j: (LEAST i. HH(f,A,i)={A}). HH(f,A,j)")] exI 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    43
by (Asm_full_simp_tac 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3731
diff changeset
    44
by (fast_tac (claset() addSIs [Ord_Least, lam_type RS domain_of_fun]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1205
diff changeset
    45
                addSEs [less_Least_subset_x, lemma1, lemma2]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents: 1123
diff changeset
    46
qed "AC15_WO6";