src/HOL/UNITY/LessThan.ML
author paulson
Thu, 11 Nov 1999 10:25:17 +0100
changeset 8005 b64d86018785
parent 7915 c7fd7eb3b0ef
child 8128 3a5864b465e2
permissions -rw-r--r--
new-style infix declaration for "image"
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     1
(*  Title:      HOL/LessThan/LessThan
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     2
    ID:         $Id$
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     5
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     6
lessThan, greaterThan, atLeast, atMost
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     7
*)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     8
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     9
7878
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    10
(*** Restrict [MOVE TO RELATION.THY??] ***)
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    11
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    12
Goalw [Restrict_def] "((x,y): Restrict A r) = ((x,y): r & x: A)";
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    13
by (Blast_tac 1);
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    14
qed "Restrict_iff";
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    15
AddIffs [Restrict_iff];
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    16
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    17
Goal "Restrict UNIV = id";
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    18
by (rtac ext 1);
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    19
by (auto_tac (claset(), simpset() addsimps [Restrict_def]));
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    20
qed "Restrict_UNIV";
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    21
Addsimps [Restrict_UNIV];
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    22
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    23
Goal "Restrict {} r = {}";
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    24
by (auto_tac (claset(), simpset() addsimps [Restrict_def]));
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    25
qed "Restrict_empty";
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    26
Addsimps [Restrict_empty];
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    27
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    28
Goalw [Restrict_def] "Restrict A (Restrict B r) = Restrict (A Int B) r";
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    29
by (Blast_tac 1);
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    30
qed "Restrict_Int";
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    31
Addsimps [Restrict_Int];
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    32
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    33
Goalw [Restrict_def] "Domain r <= A ==> Restrict A r = r";
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    34
by Auto_tac;
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    35
qed "Restrict_triv";
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    36
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    37
Goalw [Restrict_def] "Restrict A r <= r";
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    38
by Auto_tac;
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    39
qed "Restrict_subset";
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    40
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    41
Goalw [Restrict_def]
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    42
     "[| A <= B;  Restrict B r = Restrict B s |] \
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    43
\     ==> Restrict A r = Restrict A s";
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    44
by (blast_tac (claset() addSEs [equalityE]) 1);
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    45
qed "Restrict_eq_mono";
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    46
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    47
Goalw [Restrict_def, image_def]
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    48
     "[| s : RR;  Restrict A r = Restrict A s |] \
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    49
\     ==> Restrict A r : Restrict A `` RR";
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    50
by Auto_tac;
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    51
qed "Restrict_imageI";
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    52
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    53
Goal "(Restrict A `` (Restrict A `` r - s)) = (Restrict A `` r - s)";
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    54
by Auto_tac;
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    55
by (rewrite_goals_tac [image_def, Restrict_def]);
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    56
by (Blast_tac 1);
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    57
qed "Restrict_image_Diff";
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    58
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    59
(*Nothing to do with Restrict, but to specialized for Fun.ML?*)
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    60
Goal "f``(g``A - B) = (UN x: (A - g-`` B). {f(g(x))})";
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    61
by (Blast_tac 1);
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    62
qed "image_Diff_image_eq";
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    63
7915
c7fd7eb3b0ef ALMOST working version: LocalTo results commented out
paulson
parents: 7878
diff changeset
    64
Goal "Domain (Restrict A r) = A Int Domain r";
c7fd7eb3b0ef ALMOST working version: LocalTo results commented out
paulson
parents: 7878
diff changeset
    65
by (Blast_tac 1);
c7fd7eb3b0ef ALMOST working version: LocalTo results commented out
paulson
parents: 7878
diff changeset
    66
qed "Domain_Restrict";
c7fd7eb3b0ef ALMOST working version: LocalTo results commented out
paulson
parents: 7878
diff changeset
    67
c7fd7eb3b0ef ALMOST working version: LocalTo results commented out
paulson
parents: 7878
diff changeset
    68
Goal "(Restrict A r) ^^ B = r ^^ (A Int B)";
c7fd7eb3b0ef ALMOST working version: LocalTo results commented out
paulson
parents: 7878
diff changeset
    69
by (Blast_tac 1);
c7fd7eb3b0ef ALMOST working version: LocalTo results commented out
paulson
parents: 7878
diff changeset
    70
qed "Image_Restrict";
c7fd7eb3b0ef ALMOST working version: LocalTo results commented out
paulson
parents: 7878
diff changeset
    71
c7fd7eb3b0ef ALMOST working version: LocalTo results commented out
paulson
parents: 7878
diff changeset
    72
Addsimps [Domain_Restrict, Image_Restrict];
c7fd7eb3b0ef ALMOST working version: LocalTo results commented out
paulson
parents: 7878
diff changeset
    73
c7fd7eb3b0ef ALMOST working version: LocalTo results commented out
paulson
parents: 7878
diff changeset
    74
7878
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7521
diff changeset
    75
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    76
(*** lessThan ***)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    77
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    78
Goalw [lessThan_def] "(i: lessThan k) = (i<k)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    79
by (Blast_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    80
qed "lessThan_iff";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    81
AddIffs [lessThan_iff];
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    82
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    83
Goalw [lessThan_def] "lessThan 0 = {}";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    84
by (Simp_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    85
qed "lessThan_0";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    86
Addsimps [lessThan_0];
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    87
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    88
Goalw [lessThan_def] "lessThan (Suc k) = insert k (lessThan k)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    89
by (simp_tac (simpset() addsimps [less_Suc_eq]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    90
by (Blast_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    91
qed "lessThan_Suc";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    92
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5625
diff changeset
    93
Goalw [lessThan_def, atMost_def] "lessThan (Suc k) = atMost k";
fe887910e32e specifications as sets of programs
paulson
parents: 5625
diff changeset
    94
by (simp_tac (simpset() addsimps [less_Suc_eq_le]) 1);
fe887910e32e specifications as sets of programs
paulson
parents: 5625
diff changeset
    95
qed "lessThan_Suc_atMost";
fe887910e32e specifications as sets of programs
paulson
parents: 5625
diff changeset
    96
6825
30e09714eef5 new finiteness theorems
paulson
parents: 6707
diff changeset
    97
Goal "finite (lessThan k)";
30e09714eef5 new finiteness theorems
paulson
parents: 6707
diff changeset
    98
by (induct_tac "k" 1);
30e09714eef5 new finiteness theorems
paulson
parents: 6707
diff changeset
    99
by (simp_tac (simpset() addsimps [lessThan_Suc]) 2);
30e09714eef5 new finiteness theorems
paulson
parents: 6707
diff changeset
   100
by Auto_tac;
30e09714eef5 new finiteness theorems
paulson
parents: 6707
diff changeset
   101
qed "finite_lessThan";
30e09714eef5 new finiteness theorems
paulson
parents: 6707
diff changeset
   102
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   103
Goal "(UN m. lessThan m) = UNIV";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   104
by (Blast_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   105
qed "UN_lessThan_UNIV";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   106
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   107
Goalw [lessThan_def, atLeast_def, le_def]
5490
85855f65d0c6 From Compl(A) to -A
paulson
parents: 5355
diff changeset
   108
    "-lessThan k = atLeast k";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   109
by (Blast_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   110
qed "Compl_lessThan";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   111
7521
a18adacbdbd2 new theorem single_Diff_lessThan
paulson
parents: 6825
diff changeset
   112
Goal "{k} - lessThan k = {k}";
a18adacbdbd2 new theorem single_Diff_lessThan
paulson
parents: 6825
diff changeset
   113
by (Blast_tac 1);
a18adacbdbd2 new theorem single_Diff_lessThan
paulson
parents: 6825
diff changeset
   114
qed "single_Diff_lessThan";
a18adacbdbd2 new theorem single_Diff_lessThan
paulson
parents: 6825
diff changeset
   115
Addsimps [single_Diff_lessThan];
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   116
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   117
(*** greaterThan ***)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   118
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   119
Goalw [greaterThan_def] "(i: greaterThan k) = (k<i)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   120
by (Blast_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   121
qed "greaterThan_iff";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   122
AddIffs [greaterThan_iff];
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   123
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   124
Goalw [greaterThan_def] "greaterThan 0 = range Suc";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   125
by (blast_tac (claset() addIs [Suc_pred RS sym]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   126
qed "greaterThan_0";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   127
Addsimps [greaterThan_0];
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   128
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   129
Goalw [greaterThan_def] "greaterThan (Suc k) = greaterThan k - {Suc k}";
5625
77e9ab9cd7b1 polymorphic versions of nat_neq_iff and nat_neqE
paulson
parents: 5596
diff changeset
   130
by (auto_tac (claset() addEs [linorder_neqE], simpset()));
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   131
qed "greaterThan_Suc";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   132
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   133
Goal "(INT m. greaterThan m) = {}";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   134
by (Blast_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   135
qed "INT_greaterThan_UNIV";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   136
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   137
Goalw [greaterThan_def, atMost_def, le_def]
5490
85855f65d0c6 From Compl(A) to -A
paulson
parents: 5355
diff changeset
   138
    "-greaterThan k = atMost k";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   139
by (Blast_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   140
qed "Compl_greaterThan";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   141
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   142
Goalw [greaterThan_def, atMost_def, le_def]
5490
85855f65d0c6 From Compl(A) to -A
paulson
parents: 5355
diff changeset
   143
    "-atMost k = greaterThan k";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   144
by (Blast_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   145
qed "Compl_atMost";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   146
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   147
Goal "less_than ^^ {k} = greaterThan k";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   148
by (Blast_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   149
qed "Image_less_than";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   150
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   151
Addsimps [Compl_greaterThan, Compl_atMost, Image_less_than];
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   152
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   153
(*** atLeast ***)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   154
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   155
Goalw [atLeast_def] "(i: atLeast k) = (k<=i)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   156
by (Blast_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   157
qed "atLeast_iff";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   158
AddIffs [atLeast_iff];
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   159
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   160
Goalw [atLeast_def, UNIV_def] "atLeast 0 = UNIV";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   161
by (Simp_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   162
qed "atLeast_0";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   163
Addsimps [atLeast_0];
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   164
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   165
Goalw [atLeast_def] "atLeast (Suc k) = atLeast k - {k}";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   166
by (simp_tac (simpset() addsimps [Suc_le_eq]) 1);
5596
b29d18d8c4d2 abstype of programs
paulson
parents: 5490
diff changeset
   167
by (simp_tac (simpset() addsimps [order_le_less]) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   168
by (Blast_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   169
qed "atLeast_Suc";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   170
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   171
Goal "(UN m. atLeast m) = UNIV";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   172
by (Blast_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   173
qed "UN_atLeast_UNIV";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   174
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   175
Goalw [lessThan_def, atLeast_def, le_def]
5490
85855f65d0c6 From Compl(A) to -A
paulson
parents: 5355
diff changeset
   176
    "-atLeast k = lessThan k";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   177
by (Blast_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   178
qed "Compl_atLeast";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   179
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   180
Goal "less_than^-1 ^^ {k} = lessThan k";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   181
by (Blast_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   182
qed "Image_inverse_less_than";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   183
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   184
Addsimps [Compl_lessThan, Compl_atLeast, Image_inverse_less_than];
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   185
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   186
(*** atMost ***)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   187
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   188
Goalw [atMost_def] "(i: atMost k) = (i<=k)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   189
by (Blast_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   190
qed "atMost_iff";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   191
AddIffs [atMost_iff];
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   192
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   193
Goalw [atMost_def] "atMost 0 = {0}";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   194
by (Simp_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   195
qed "atMost_0";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   196
Addsimps [atMost_0];
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   197
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   198
Goalw [atMost_def] "atMost (Suc k) = insert (Suc k) (atMost k)";
5596
b29d18d8c4d2 abstype of programs
paulson
parents: 5490
diff changeset
   199
by (simp_tac (simpset() addsimps [less_Suc_eq, order_le_less]) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   200
by (Blast_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   201
qed "atMost_Suc";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   202
6825
30e09714eef5 new finiteness theorems
paulson
parents: 6707
diff changeset
   203
Goal "finite (atMost k)";
30e09714eef5 new finiteness theorems
paulson
parents: 6707
diff changeset
   204
by (induct_tac "k" 1);
30e09714eef5 new finiteness theorems
paulson
parents: 6707
diff changeset
   205
by (simp_tac (simpset() addsimps [atMost_Suc]) 2);
30e09714eef5 new finiteness theorems
paulson
parents: 6707
diff changeset
   206
by Auto_tac;
30e09714eef5 new finiteness theorems
paulson
parents: 6707
diff changeset
   207
qed "finite_atMost";
30e09714eef5 new finiteness theorems
paulson
parents: 6707
diff changeset
   208
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   209
Goal "(UN m. atMost m) = UNIV";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   210
by (Blast_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   211
qed "UN_atMost_UNIV";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   212
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   213
Goalw [atMost_def, le_def]
5490
85855f65d0c6 From Compl(A) to -A
paulson
parents: 5355
diff changeset
   214
    "-atMost k = greaterThan k";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   215
by (Blast_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   216
qed "Compl_atMost";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   217
Addsimps [Compl_atMost];
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   218
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   219
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   220
(*** Combined properties ***)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   221
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   222
Goal "atMost n Int atLeast n = {n}";
6707
3b07e62a718c using generic rules when possible
paulson
parents: 6128
diff changeset
   223
by (blast_tac (claset() addIs [order_antisym]) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   224
qed "atMost_Int_atLeast";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   225
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   226
5232
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   227
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   228
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   229
(*** Finally, a few set-theoretic laws...
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   230
     CAN BOOLEAN SIMPLIFICATION BE AUTOMATED? ***)
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   231
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   232
context Set.thy;
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   233
5490
85855f65d0c6 From Compl(A) to -A
paulson
parents: 5355
diff changeset
   234
(** Rewrite rules to eliminate A.  Conditions can be satisfied by letting B
5232
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   235
    be any set including A Int C and contained in A Un C, such as B=A or B=C.
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   236
**)
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   237
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   238
Goal "[| A Int C <= B; B <= A Un C |] \
5490
85855f65d0c6 From Compl(A) to -A
paulson
parents: 5355
diff changeset
   239
\     ==> (A Int B) Un (-A Int C) = B Un C";
5232
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   240
by (Blast_tac 1);
5490
85855f65d0c6 From Compl(A) to -A
paulson
parents: 5355
diff changeset
   241
qed "set_cancel1";
5232
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   242
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   243
Goal "[| A Int C <= B; B <= A Un C |] \
5490
85855f65d0c6 From Compl(A) to -A
paulson
parents: 5355
diff changeset
   244
\     ==> (A Un B) Int (-A Un C) = B Int C";
5232
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   245
by (Blast_tac 1);
5490
85855f65d0c6 From Compl(A) to -A
paulson
parents: 5355
diff changeset
   246
qed "set_cancel2";
5232
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   247
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   248
(*The base B=A*)
5490
85855f65d0c6 From Compl(A) to -A
paulson
parents: 5355
diff changeset
   249
Goal "A Un (-A Int C) = A Un C";
5232
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   250
by (Blast_tac 1);
5490
85855f65d0c6 From Compl(A) to -A
paulson
parents: 5355
diff changeset
   251
qed "set_cancel3";
5232
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   252
5490
85855f65d0c6 From Compl(A) to -A
paulson
parents: 5355
diff changeset
   253
Goal "A Int (-A Un C) = A Int C";
5232
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   254
by (Blast_tac 1);
5490
85855f65d0c6 From Compl(A) to -A
paulson
parents: 5355
diff changeset
   255
qed "set_cancel4";
5232
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   256
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   257
(*The base B=C*)
5490
85855f65d0c6 From Compl(A) to -A
paulson
parents: 5355
diff changeset
   258
Goal "(A Int C) Un (-A Int C) = C";
5232
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   259
by (Blast_tac 1);
5490
85855f65d0c6 From Compl(A) to -A
paulson
parents: 5355
diff changeset
   260
qed "set_cancel5";
5232
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   261
5490
85855f65d0c6 From Compl(A) to -A
paulson
parents: 5355
diff changeset
   262
Goal "(A Un C) Int (-A Un C) = C";
5232
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   263
by (Blast_tac 1);
5490
85855f65d0c6 From Compl(A) to -A
paulson
parents: 5355
diff changeset
   264
qed "set_cancel6";
85855f65d0c6 From Compl(A) to -A
paulson
parents: 5355
diff changeset
   265
85855f65d0c6 From Compl(A) to -A
paulson
parents: 5355
diff changeset
   266
Addsimps [set_cancel1, set_cancel2, set_cancel3,
85855f65d0c6 From Compl(A) to -A
paulson
parents: 5355
diff changeset
   267
	  set_cancel4, set_cancel5, set_cancel6];
5232
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   268
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   269
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   270
(** More ad-hoc rules **)
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   271
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   272
Goal "A Un B - (A - B) = B";
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   273
by (Blast_tac 1);
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   274
qed "Un_Diff_Diff";
5490
85855f65d0c6 From Compl(A) to -A
paulson
parents: 5355
diff changeset
   275
Addsimps [Un_Diff_Diff];
5232
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   276
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   277
Goal "A Int (B - C) Un C = A Int B Un C";
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   278
by (Blast_tac 1);
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   279
qed "Int_Diff_Un";
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   280
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   281
Goal "Union(B) Int A = (UN C:B. C Int A)";
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   282
by (Blast_tac 1);
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   283
qed "Int_Union2";
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   284
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   285
Goal "Union(B) Int A = Union((%C. C Int A)``B)";
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   286
by (Blast_tac 1);
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   287
qed "Int_Union_Union";
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5069
diff changeset
   288