diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/UNITY/Comp.thy --- a/src/ZF/UNITY/Comp.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/UNITY/Comp.thy Tue Sep 27 17:46:52 2022 +0100 @@ -18,33 +18,33 @@ theory Comp imports Union Increasing begin definition - component :: "[i,i]=>o" (infixl \component\ 65) where + component :: "[i,i]\o" (infixl \component\ 65) where "F component H \ (\G. F \ G = H)" definition - strict_component :: "[i,i]=>o" (infixl \strict'_component\ 65) where + strict_component :: "[i,i]\o" (infixl \strict'_component\ 65) where "F strict_component H \ F component H \ F\H" definition (* A stronger form of the component relation *) - component_of :: "[i,i]=>o" (infixl \component'_of\ 65) where + component_of :: "[i,i]\o" (infixl \component'_of\ 65) where "F component_of H \ (\G. F ok G \ F \ G = H)" definition - strict_component_of :: "[i,i]=>o" (infixl \strict'_component'_of\ 65) where + strict_component_of :: "[i,i]\o" (infixl \strict'_component'_of\ 65) where "F strict_component_of H \ F component_of H \ F\H" definition (* Preserves a state function f, in particular a variable *) - preserves :: "(i=>i)=>i" where + preserves :: "(i\i)\i" where "preserves(f) \ {F:program. \z. F: stable({s:state. f(s) = z})}" definition - fun_pair :: "[i=>i, i =>i] =>(i=>i)" where - "fun_pair(f, g) \ %x. " + fun_pair :: "[i\i, i \i] \(i\i)" where + "fun_pair(f, g) \ \x. " definition - localize :: "[i=>i, i] => i" where + localize :: "[i\i, i] \ i" where "localize(f,F) \ mk_program(Init(F), Acts(F), AllowedActs(F) \ (\G\preserves(f). Acts(G)))" @@ -156,7 +156,7 @@ done lemma preserves_imp_eq: - "\F \ preserves(f); act \ Acts(F); \ act\ \ f(s) = f(t)" + "\F \ preserves(f); act \ Acts(F); \s,t\ \ act\ \ f(s) = f(t)" apply (unfold preserves_def stable_def constrains_def) apply (subgoal_tac "s \ state \ t \ state") prefer 2 apply (blast dest!: Acts_type [THEN subsetD], force) @@ -226,11 +226,11 @@ "\F \ preserves(f); \x \ state. f(x):A\ \ F \ Increasing.increasing(A, r, f)" apply (unfold Increasing.increasing_def) apply (auto intro: preserves_into_program) -apply (rule_tac P = "%x. :r" in preserves_imp_stable, auto) +apply (rule_tac P = "\x. \k, x\:r" in preserves_imp_stable, auto) done lemma preserves_id_subset_stable: - "st_set(A) \ preserves(%x. x) \ stable(A)" + "st_set(A) \ preserves(\x. x) \ stable(A)" apply (unfold preserves_def stable_def constrains_def, auto) apply (drule_tac x = xb in spec) apply (drule_tac x = act in bspec) @@ -238,7 +238,7 @@ done lemma preserves_id_imp_stable: - "\F \ preserves(%x. x); st_set(A)\ \ F \ stable(A)" + "\F \ preserves(\x. x); st_set(A)\ \ F \ stable(A)" by (blast intro: preserves_id_subset_stable [THEN subsetD]) (** Added by Sidi **)