include "basics/types.ma". include "ASM/String.ma". include "ASM/Arithmetic.ma". include "common/Errors.ma". (* identifiers and their generators are tagged to differentiate them, and to provide extra type checking. *) inductive identifier (tag:String) : Type[0] ≝ an_identifier : Word → identifier tag. record universe (tag:String) : Type[0] ≝ { next_identifier : Word }. definition new_universe : ∀tag:String. universe tag ≝ λtag. mk_universe tag (zero ?). definition fresh : ∀tag. universe tag → res (identifier tag × (universe tag)) ≝ λtag,g. let 〈gen, carries〉 ≝ add_with_carries ? (next_identifier ? g) (zero ?) true in if get_index_v ?? carries 0 ? then Error ? else OK ? 〈an_identifier tag (next_identifier ? g), mk_universe tag gen〉. // qed. definition identifier_eq : ∀tag:String. ∀x,y:identifier tag. (x=y) + (x≠y). #tag * #x * #y lapply (refl ? (eq_bv ? x y)) cases (eq_bv ? x y) in ⊢ (???% → %) #E [ % | %2 ] lapply E @eq_bv_elim [ #H #_ >H @refl | 2,3: #_ #H destruct | #H #_ % #H' destruct /2/ ] qed. definition identifier_of_nat : ∀tag:String. nat → identifier tag ≝ λtag,n. an_identifier tag (bitvector_of_nat ? n). (* Maps from identifiers to arbitrary types. *) include "ASM/BitVectorTrie.ma". inductive identifier_map (tag:String) (A:Type[0]) : Type[0] ≝ an_id_map : BitVectorTrie A 16 → identifier_map tag A. definition empty_map : ∀tag:String. ∀A. identifier_map tag A ≝ λtag,A. an_id_map tag A (Stub A 16). definition lookup : ∀tag,A. identifier_map tag A → identifier tag → option A ≝ λtag,A,m,l. lookup_opt A 16 (match l with [ an_identifier l' ⇒ l' ]) (match m with [ an_id_map m' ⇒ m' ]). definition add : ∀tag,A. identifier_map tag A → identifier tag → A → identifier_map tag A ≝ λtag,A,m,l,a. an_id_map tag A (insert A 16 (match l with [ an_identifier l' ⇒ l' ]) a (match m with [ an_id_map m' ⇒ m' ])).