diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/UNITY/State.thy --- a/src/ZF/UNITY/State.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/UNITY/State.thy Tue Sep 27 17:46:52 2022 +0100 @@ -17,8 +17,8 @@ type_intros nat_subset_univ [THEN list_subset_univ, THEN subsetD] consts - type_of :: "i=>i" - default_val :: "i=>i" + type_of :: "i\i" + default_val :: "i\i" definition "state \ \x \ var. cons(default_val(x), type_of(x))" @@ -27,13 +27,13 @@ "st0 \ \x \ var. default_val(x)" definition - st_set :: "i=>o" where + st_set :: "i\o" where (* To prevent typing conditions like `A<=state' from being used in combination with the rules `constrains_weaken', etc. *) "st_set(A) \ A<=state" definition - st_compl :: "i=>i" where + st_compl :: "i\i" where "st_compl(A) \ state-A"