| author | wenzelm | 
| Sat, 02 Sep 2000 21:49:51 +0200 | |
| changeset 9803 | bc883b390d91 | 
| parent 9248 | e1dee89de037 | 
| 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 | |
| 5068 | 16 | Goalw [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 | |
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
5068diff
changeset | 20 | Goalw [less_lift_def] | 
| 3323 
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
 slotosch parents: 
3033diff
changeset | 21 | "[|(x1::'a lift) << x2; x2 << x1|] ==> x1 = x2"; | 
| 2356 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 22 | by (fast_tac HOL_cs 1); | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 23 | qed"antisym_less_lift"; | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 24 | |
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
5068diff
changeset | 25 | Goalw [less_lift_def] | 
| 3323 
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
 slotosch parents: 
3033diff
changeset | 26 | "[|(x1::'a lift) << x2; x2 << x3|] ==> x1 << x3"; | 
| 2356 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 27 | by (fast_tac HOL_cs 1); | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 28 | qed"trans_less_lift"; |