author | nipkow |
Thu, 26 Jun 1997 10:42:50 +0200 | |
changeset 3460 | 5d71eed16fbe |
parent 3323 | 194ae2e0c193 |
child 5068 | fb28eaa07e01 |
permissions | -rw-r--r-- |
2357 | 1 |
(* Title: HOLCF/Lift1.ML |
2 |
ID: $Id$ |
|
3033 | 3 |
Author: Olaf Mueller |
2357 | 4 |
Copyright 1996 Technische Universitaet Muenchen |
5 |
||
6 |
Theorems for Lift1.thy |
|
7 |
*) |
|
8 |
||
2356
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
9 |
|
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
10 |
open Lift1; |
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
11 |
|
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
12 |
(* ------------------------------------------------------------------------ *) |
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
13 |
(* less_lift is a partial order on type 'a -> 'b *) |
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
14 |
(* ------------------------------------------------------------------------ *) |
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
15 |
|
3323
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
slotosch
parents:
3033
diff
changeset
|
16 |
goalw thy [less_lift_def] "(x::'a lift) << x"; |
2356
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
17 |
by (fast_tac HOL_cs 1); |
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
18 |
qed"refl_less_lift"; |
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
19 |
|
2640 | 20 |
val prems = goalw thy [less_lift_def] |
3323
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
slotosch
parents:
3033
diff
changeset
|
21 |
"[|(x1::'a lift) << x2; x2 << x1|] ==> x1 = x2"; |
2640 | 22 |
by (cut_facts_tac prems 1); |
2356
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
23 |
by (fast_tac HOL_cs 1); |
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
24 |
qed"antisym_less_lift"; |
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
25 |
|
2640 | 26 |
val prems = goalw Lift1.thy [less_lift_def] |
3323
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
slotosch
parents:
3033
diff
changeset
|
27 |
"[|(x1::'a lift) << x2; x2 << x3|] ==> x1 << x3"; |
2640 | 28 |
by (cut_facts_tac prems 1); |
2356
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
29 |
by (fast_tac HOL_cs 1); |
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff
changeset
|
30 |
qed"trans_less_lift"; |