src/HOLCF/Porder0.thy
author slotosch
Mon, 17 Feb 1997 10:57:11 +0100
changeset 2640 ee4dfce170a0
parent 2394 91d8abf108be
child 2850 a66196e1668c
permissions -rw-r--r--
Changes of HOLCF from Oscar Slotosch: 1. axclass instead of class * less instead of less_fun, less_cfun, less_sprod, less_cprod, less_ssum, less_up, less_lift * @x.!y.x<<y instead of UUU instead of UU_fun, UU_cfun, ... * no witness type void needed (eliminated Void.thy.Void.ML) * inst_<typ>_<class> derived as theorems 2. improved some proves on less_sprod and less_cprod * eliminated the following theorems Sprod1.ML: less_sprod1a Sprod1.ML: less_sprod1b Sprod1.ML: less_sprod2a Sprod1.ML: less_sprod2b Sprod1.ML: less_sprod2c Sprod2.ML: less_sprod3a Sprod2.ML: less_sprod3b Sprod2.ML: less_sprod4b Sprod2.ML: less_sprod4c Sprod3.ML: less_sprod5b Sprod3.ML: less_sprod5c Cprod1.ML: less_cprod1b Cprod1.ML: less_cprod2a Cprod1.ML: less_cprod2b Cprod1.ML: less_cprod2c Cprod2.ML: less_cprod3a Cprod2.ML: less_cprod3b 3. new classes: * cpo<po, * chfin<pcpo, * flat<pcpo, * derived: flat<chfin to do: show instances for lift 4. Data Type One * Used lift for the definition: one = unit lift * Changed the constant one into ONE 5. Data Type Tr * Used lift for the definition: tr = bool lift * adopted definitions of if,andalso,orelse,neg * only one theory Tr.thy,Tr.ML instead of Tr1.thy,Tr1.ML, Tr2.thy,Tr2.ML * reintroduced ceils for =TT,=FF 6. typedef * Using typedef instead of faking type definitions to do: change fapp, fabs from Cfun1 to Rep_Cfun, Abs_Cfun 7. adopted examples and domain construct to theses changes These changes eliminated all rules and arities from HOLCF
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2394
diff changeset
     1
(*  Title:      HOLCF/Porder0.thy
298
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
     2
    ID:         $Id$
1479
21eb5e156d91 expanded tabs
clasohm
parents: 1274
diff changeset
     3
    Author:     Franz Regensburger
298
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
     4
    Copyright   1993 Technische Universitaet Muenchen
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
     5
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
     6
Definition of class porder (partial order)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
     7
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
     8
*)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
     9
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2394
diff changeset
    10
Porder0 = Nat +
298
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    11
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2394
diff changeset
    12
(* first the global constant for HOLCF type classes *)
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2394
diff changeset
    13
consts
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2394
diff changeset
    14
  "less"        :: "['a,'a] => bool" (infixl "\\<sqsubseteq>\\<sqsubseteq>" 55)
298
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    15
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2394
diff changeset
    16
axclass po < term
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2394
diff changeset
    17
        (* class axioms: *)
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2394
diff changeset
    18
ax_refl_less       "less x x"        
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2394
diff changeset
    19
ax_antisym_less    "[|less x y; less y x |] ==> x = y"    
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2394
diff changeset
    20
ax_trans_less      "[|less x y; less y z |] ==> less x z"
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2394
diff changeset
    21
 
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2394
diff changeset
    22
	(* characteristic constant << on po *)
2394
91d8abf108be adaptions for symbol font
oheimb
parents: 2291
diff changeset
    23
consts
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2394
diff changeset
    24
  "<<"          :: "['a,'a::po] => bool"        (infixl 55)
2394
91d8abf108be adaptions for symbol font
oheimb
parents: 2291
diff changeset
    25
91d8abf108be adaptions for symbol font
oheimb
parents: 2291
diff changeset
    26
syntax (symbols)
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2394
diff changeset
    27
  "op <<"       :: "['a,'a::po] => bool"        (infixl "\\<sqsubseteq>" 55)
298
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    28
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2394
diff changeset
    29
defs
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2394
diff changeset
    30
po_def             "(op <<) == less"
298
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    31
end 
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents: 442
diff changeset
    32
ea0668a1c0ba added 8bit pragmas
regensbu
parents: 442
diff changeset
    33