src/ZF/UNITY/UNITYMisc.ML
author ehmety
Thu, 15 Nov 2001 20:01:19 +0100
changeset 12215 55c084921240
parent 11479 697dcaaf478f
child 14046 6616e6c53d48
permissions -rw-r--r--
Modified to make the files build with the new changes in ZF
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/UNITYMisc.ML
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     2
    ID:         $Id$
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     3
    Author:     Sidi O Ehmety, Computer Laboratory
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     4
    Copyright   2001  University of Cambridge
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     5
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     6
Some miscellaneous and add-hoc set theory concepts.
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     7
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     8
*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     9
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    10
Goalw [measure_def, less_than_def] 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    11
         "less_than(A) = {<x,y>:A*A. x<y}";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    12
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    13
qed "less_than_equals";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    14
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    15
Goalw [less_than_def] "wf(less_than(A))";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    16
by (rtac wf_measure 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    17
qed "wf_less_than";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    18
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    19
Goalw [less_than_def, measure_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    20
       "less_than(A)<= A*A";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    21
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    22
qed "less_than_subset";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    23
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    24
Goalw [less_than_def, measure_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    25
"<x,y>:less_than(A) <-> (x:A & y:A & x<y)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    26
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    27
qed "less_than_iff";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    28
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    29
Goalw [lessThan_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    30
"i:lessThan(k,A) <-> i:A & i<k";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    31
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    32
qed "lessThan_iff";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    33
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    34
Goalw [greaterThan_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    35
"i:greaterThan(k,A) <-> i:A & k<i";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    36
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    37
qed "greaterThan_iff";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    38
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    39
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    40
(** Needed for WF reasoning in WFair.ML **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    41
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    42
Goal "k:A ==>less_than(A)``{k} = greaterThan(k, A)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    43
by (rtac equalityI 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    44
by (auto_tac (claset(), simpset() addsimps 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    45
                [less_than_iff,greaterThan_iff]));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    46
qed "Image_less_than";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    47
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    48
Goal "k:A ==> less_than(A)-`` {k} = lessThan(k, A)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    49
by (rtac equalityI 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    50
by (auto_tac (claset(), simpset() addsimps 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    51
                [less_than_iff,lessThan_iff]));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    52
qed "Image_inverse_less_than";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    53
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    54
Addsimps [Image_less_than, Image_inverse_less_than];
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    55
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    56
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    57
(** Ad-hoc set-theory rules **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    58
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    59
Goal "Union(B) Int A = (UN b:B. b Int A)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    60
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    61
qed "Int_Union_Union";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    62
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    63
Goal "A Int Union(B) = (UN b:B. A Int b)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    64
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    65
qed "Int_Union_Union2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    66
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    67
Goal "A Un B - (A - B) = B";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    68
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    69
qed "Un_Diff_Diff";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    70
AddIffs [Un_Diff_Diff];
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    71
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    72
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    73
(** Needed in State theory for the current definition of variables 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    74
where they are indexed by lists  **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    75
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    76
Goal "i:list(nat) ==> i:univ(0)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    77
by (dres_inst_tac [("B", "0")] list_into_univ 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    78
by (blast_tac (claset() addIs [nat_into_univ]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    79
by (assume_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    80
qed "list_nat_into_univ";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    81
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    82
(** To be moved to Update theory **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    83
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    84
Goalw [update_def] 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    85
  "[| f:Pi(A,B); x:A; y:B(x) |] ==> f(x:=y):Pi(A, B)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    86
by (asm_simp_tac (simpset() addsimps [domain_of_fun, cons_absorb, 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    87
                                      apply_funtype, lam_type]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    88
qed "update_type2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    89
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    90
(** Simplication rules for Collect; To be moved elsewhere **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    91
Goal "{x:A. P(x)} Int A = {x:A. P(x)}";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    92
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    93
qed "Collect_Int2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    94
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    95
Goal "A Int {x:A. P(x)} = {x:A. P(x)}";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    96
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    97
qed "Collect_Int3";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    98
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    99
AddIffs [Collect_Int2, Collect_Int3];
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   100
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   101
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   102
Goal "{x:A. P(x) | Q(x)} = Collect(A, P) Un Collect(A, Q)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   103
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   104
qed "Collect_disj_eq";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   105
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   106
Goal "{x:A. P(x) & Q(x)} = Collect(A, P) Int Collect(A, Q)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   107
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   108
qed "Collect_conj_eq";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   109
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   110
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   111