{-# OPTIONS --allow-unsolved-metas #-} open import Level hiding (suc ; zero ) open import Ordinals module cardinal {n : Level } (O : Ordinals {n}) where open import logic -- import OD import OD hiding ( _⊆_ ) import ODC open import Data.Nat open import Relation.Binary.PropositionalEquality open import Data.Nat.Properties open import Data.Empty open import Relation.Nullary open import Relation.Binary open import Relation.Binary.Core open inOrdinal O open OD O open OD.OD open ODAxiom odAxiom open import ZProduct O import OrdUtil import ODUtil open Ordinals.Ordinals O open Ordinals.IsOrdinals isOrdinal -- open Ordinals.IsNext isNext open OrdUtil O open ODUtil O _⊆_ : ( A B : HOD ) → Set n _⊆_ A B = {x : Ordinal } → odef A x → odef B x open _∧_ open _∨_ open Bool open _==_ open HOD record OrdBijection (A B : Ordinal ) : Set n where field fun← : (x : Ordinal ) → odef (* A) x → Ordinal fun→ : (x : Ordinal ) → odef (* B) x → Ordinal funB : (x : Ordinal ) → ( lt : odef (* A) x ) → odef (* B) ( fun← x lt ) funA : (x : Ordinal ) → ( lt : odef (* B) x ) → odef (* A) ( fun→ x lt ) fiso← : (x : Ordinal ) → ( lt : odef (* B) x ) → fun← ( fun→ x lt ) ( funA x lt ) ≡ x fiso→ : (x : Ordinal ) → ( lt : odef (* A) x ) → fun→ ( fun← x lt ) ( funB x lt ) ≡ x ordbij-refl : { a b : Ordinal } → a ≡ b → OrdBijection a b ordbij-refl {a} refl = record { fun← = λ x _ → x ; fun→ = λ x _ → x ; funB = λ x lt → lt ; funA = λ x lt → lt ; fiso← = λ x lt → refl ; fiso→ = λ x lt → refl } open Injection open OrdBijection record IsImage (a b : Ordinal) (iab : Injection a b ) (x : Ordinal ) : Set n where field y : Ordinal ay : odef (* a) y x=fy : x ≡ i→ iab _ ay Image : (a : Ordinal) { b : Ordinal } → Injection a b → HOD Image a {b} iab = record { od = record { def = λ x → IsImage a b iab x } ; odmax = b ;