diff -r 5b946c81dfbf -r 324bc1ffba12 src/HOL/TLA/Action.thy --- a/src/HOL/TLA/Action.thy Mon Jan 11 21:21:02 2016 +0100 +++ b/src/HOL/TLA/Action.thy Mon Jan 11 22:23:03 2016 +0100 @@ -9,25 +9,19 @@ imports Stfun begin - -(** abstract syntax **) - -type_synonym 'a trfun = "(state * state) \ 'a" -type_synonym action = "bool trfun" +type_synonym 'a trfun = "state \ state \ 'a" +type_synonym action = "bool trfun" instance prod :: (world, world) world .. -consts - (** abstract syntax **) - before :: "'a stfun \ 'a trfun" - after :: "'a stfun \ 'a trfun" - unch :: "'a stfun \ action" +definition enabled :: "action \ stpred" + where "enabled A s \ \u. (s,u) \ A" + - SqAct :: "[action, 'a stfun] \ action" - AnAct :: "[action, 'a stfun] \ action" - enabled :: "action \ stpred" - -(** concrete syntax **) +consts + before :: "'a stfun \ 'a trfun" + after :: "'a stfun \ 'a trfun" + unch :: "'a stfun \ action" syntax (* Syntax for writing action expressions in arbitrary contexts *) @@ -40,8 +34,6 @@ (*** Priming: same as "after" ***) "_prime" :: "lift \ lift" ("(_`)" [100] 99) - "_SqAct" :: "[lift, lift] \ lift" ("([_]'_(_))" [0,1000] 99) - "_AnAct" :: "[lift, lift] \ lift" ("(<_>'_(_))" [0,1000] 99) "_Enabled" :: "lift \ lift" ("(Enabled _)" [100] 100) translations @@ -50,25 +42,30 @@ "_after" == "CONST after" "_prime" => "_after" "_unchanged" == "CONST unch" - "_SqAct" == "CONST SqAct" - "_AnAct" == "CONST AnAct" "_Enabled" == "CONST enabled" - "w \ [A]_v" <= "_SqAct A v w" - "w \ _v" <= "_AnAct A v w" "s \ Enabled A" <= "_Enabled A s" "w \ unchanged f" <= "_unchanged f w" axiomatization where unl_before: "(ACT $v) (s,t) \ v s" and unl_after: "(ACT v$) (s,t) \ v t" and - unchanged_def: "(s,t) \ unchanged v \ (v t = v s)" -defs - square_def: "ACT [A]_v \ ACT (A \ unchanged v)" - angle_def: "ACT _v \ ACT (A \ \ unchanged v)" + +definition SqAct :: "[action, 'a stfun] \ action" + where square_def: "SqAct A v \ ACT (A \ unchanged v)" + +definition AnAct :: "[action, 'a stfun] \ action" + where angle_def: "AnAct A v \ ACT (A \ \ unchanged v)" - enabled_def: "s \ Enabled A \ \u. (s,u) \ A" +syntax + "_SqAct" :: "[lift, lift] \ lift" ("([_]'_(_))" [0, 1000] 99) + "_AnAct" :: "[lift, lift] \ lift" ("(<_>'_(_))" [0, 1000] 99) +translations + "_SqAct" == "CONST SqAct" + "_AnAct" == "CONST AnAct" + "w \ [A]_v" \ "_SqAct A v w" + "w \ _v" \ "_AnAct A v w" (* The following assertion specializes "intI" for any world type