| author | obua | 
| Mon, 14 Jun 2004 14:20:55 +0200 | |
| changeset 14940 | b9ab8babd8b3 | 
| parent 12030 | 46d57d0290a2 | 
| child 14981 | e73f8140af78 | 
| permissions | -rw-r--r-- | 
| 2640 | 1 | (* Title: HOLCF/Cprod1.thy | 
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 2 | ID: $Id$ | 
| 1479 | 3 | Author: Franz Regensburger | 
| 12030 | 4 | License: GPL (GNU GENERAL PUBLIC LICENSE) | 
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 5 | |
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 6 | Partial ordering for cartesian product of HOL theory prod.thy | 
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 7 | *) | 
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 8 | |
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 9 | Cprod1 = Cfun3 + | 
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 10 | |
| 2840 
7e03e61612b0
generalized theorems and class instances for Cprod.
 slotosch parents: 
2640diff
changeset | 11 | default cpo | 
| 
7e03e61612b0
generalized theorems and class instances for Cprod.
 slotosch parents: 
2640diff
changeset | 12 | |
| 3323 
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
 slotosch parents: 
2840diff
changeset | 13 | instance "*"::(sq_ord,sq_ord)sq_ord | 
| 
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
 slotosch parents: 
2840diff
changeset | 14 | |
| 1168 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 regensbu parents: 
1150diff
changeset | 15 | defs | 
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 16 | |
| 3323 
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
 slotosch parents: 
2840diff
changeset | 17 | less_cprod_def "p1 << p2 == (fst p1<<fst p2 & snd p1 << snd p2)" | 
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 18 | |
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 19 | end |