chapter ‹Generated by Lem from ‹basic_classes.lem›.› theory "Lem_basic_classes" imports Main "Lem_bool" begin ― ‹‹****************************************************************************›› ― ‹‹ Basic Type Classes ›› ― ‹‹****************************************************************************›› ― ‹‹open import Bool›› ― ‹‹open import {coq} `Coq.Strings.Ascii`›› ― ‹‹open import {hol} `ternaryComparisonsTheory`›› ― ‹‹ ========================================================================== ›› ― ‹‹ Equality ›› ― ‹‹ ========================================================================== ›› ― ‹‹ Lem`s default equality (=) is defined by the following type-class Eq. This typeclass should define equality on an abstract datatype 'a. It should always coincide with the default equality of Coq, HOL and Isabelle. For OCaml, it might be different, since abstract datatypes like sets might have fancy equalities. ›› ― ‹‹class ( Eq 'a ) val = [isEqual] : 'a -> 'a -> bool val <> [isInequal] : 'a -> 'a -> bool end›› ― ‹‹ (=) should for all instances be an equivalence relation The isEquivalence predicate of relations could be used here. However, this would lead to a cyclic dependency. ›› ― ‹‹ TODO: add later, once lemmata can be assigned to classes lemma eq_equiv: ((forall x. (x = x)) && (forall x y. (x = y) <-> (y = x)) && (forall x y z. ((x = y) && (y = z)) --> (x = z))) ›› ― ‹‹ Structural equality ›› ― ‹‹ Sometimes, it is also handy to be able to use structural equality. This equality is mapped to the build-in equality of backends. This equality differs significantly for each backend. For example, OCaml can`t check equality of function types, whereas HOL can. When using structural equality, one should know what one is doing. The only guarentee is that is behaves like the native backend equality. A lengthy name for structural equality is used to discourage its direct use. It also ensures that users realise it is unsafe (e.g. OCaml can`t check two functions for equality ›› ― ‹‹val unsafe_structural_equality : forall 'a. 'a -> 'a -> bool›› ― ‹‹val unsafe_structural_inequality : forall 'a. 'a -> 'a -> bool›› ― ‹‹let unsafe_structural_inequality x y= not (unsafe_structural_equality x y)›› ― ‹‹ ========================================================================== ›› ― ‹‹ Orderings ›› ― ‹‹ ========================================================================== ›› ― ‹‹ The type-class Ord represents total orders (also called linear orders) ››