1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 include "ground_2/arith.ma".
17 (* SORT HIERARCHY ***********************************************************)
19 (* sort hierarchy specification *)
20 record sh: Type[0] ≝ {
21 next : nat → nat; (* next sort in the hierarchy *)
22 next_lt: ∀k. k < next k (* strict monotonicity condition *)
25 definition sh_N: sh ≝ mk_sh S ….
28 (* Basic properties *********************************************************)
30 lemma nexts_le: ∀h,k,l. k ≤ (next h)^l k.
31 #h #k #l elim l -l // normalize #l #IHl
32 lapply (next_lt h ((next h)^l k)) #H
33 lapply (le_to_lt_to_lt … IHl H) -IHl -H /2 width=2/
36 lemma nexts_lt: ∀h,k,l. k < (next h)^(l+1) k.
38 lapply (nexts_le h k l) #H
39 @(le_to_lt_to_lt … H) //
42 axiom nexts_dec: ∀h,k1,k2. Decidable (∃l. (next h)^l k1 = k2).
44 axiom nexts_inj: ∀h,k,l1,l2. (next h)^l1 k = (next h)^l2 k → l1 = l2.