Up to index of Isabelle/HOL/HOL-Complex/CSP/CSP_T/CSP_F
theory CSP_F_law_basic (*-------------------------------------------*
| CSP-Prover on Isabelle2004 |
| December 2004 |
| June 2005 (modified) |
| September 2005 (modified) |
| |
| CSP-Prover on Isabelle2005 |
| October 2005 (modified) |
| January 2006 (modified) |
| April 2006 (modified) |
| March 2007 (modified) |
| August 2007 (modified) |
| |
| Yoshinao Isobe (AIST JAPAN) |
*-------------------------------------------*)
theory CSP_F_law_basic
imports CSP_F_law_decompo CSP_T_law_basic
begin
(*****************************************************************
1. Commutativity
2. Associativity
3. Idempotence
4. Left Commutativity
5. IF
*****************************************************************)
(*********************************************************
IF bool
*********************************************************)
(*------------------*
| csp law |
*------------------*)
lemma cspF_IF_split:
"IF b THEN P ELSE Q =F[M,M] (if b then P else Q)"
apply (simp add: cspF_semantics)
apply (simp add: traces_def)
apply (simp add: failures_def)
done
lemma cspF_IF_True:
"IF True THEN P ELSE Q =F[M,M] P"
apply (rule cspF_rw_left)
apply (rule cspF_IF_split)
by (simp)
lemma cspF_IF_False:
"IF False THEN P ELSE Q =F[M,M] Q"
apply (rule cspF_rw_left)
apply (rule cspF_IF_split)
by (simp)
lemmas cspF_IF = cspF_IF_True cspF_IF_False
(*-----------------------------------*
| Idempotence |
*-----------------------------------*)
lemma cspF_Ext_choice_idem:
"P [+] P =F[M,M] P"
apply (simp add: cspF_cspT_semantics)
apply (simp add: cspT_Ext_choice_idem)
apply (rule order_antisym)
apply (rule, simp add: in_traces in_failures)
apply (elim conjE disjE)
apply (simp_all)
apply (rule proc_F2_F4)
apply (simp_all)
apply (rule, simp add: in_traces in_failures)
done
lemma cspF_Int_choice_idem:
"P |~| P =F[M,M] P"
apply (simp add: cspF_cspT_semantics)
apply (simp add: cspT_Int_choice_idem)
apply (rule order_antisym)
apply (rule, simp add: in_failures)+
done
(*------------------*
| csp law |
*------------------*)
lemmas cspF_idem = cspF_Ext_choice_idem cspF_Int_choice_idem
(*-----------------------------------*
| Commutativity |
*-----------------------------------*)
(*********************************************************
Ext choice
*********************************************************)
lemma cspF_Ext_choice_commut:
"P [+] Q =F[M,M] Q [+] P"
apply (simp add: cspF_cspT_semantics)
apply (simp add: cspT_Ext_choice_commut)
apply (rule order_antisym)
apply (rule, simp add: in_failures, fast)+
done
(*********************************************************
Int choice
*********************************************************)
lemma cspF_Int_choice_commut:
"P |~| Q =F[M,M] Q |~| P"
apply (simp add: cspF_cspT_semantics)
apply (simp add: cspT_Int_choice_commut)
apply (rule order_antisym)
apply (rule, simp add: in_failures, fast)+
done
(*********************************************************
Parallel
*********************************************************)
lemma cspF_Parallel_commut:
"P |[X]| Q =F[M,M] Q |[X]| P"
apply (simp add: cspF_cspT_semantics)
apply (simp add: cspT_Parallel_commut)
apply (rule order_antisym)
apply (rule, simp add: in_failures)
apply (elim conjE exE)
apply (rule_tac x="Z" in exI)
apply (rule_tac x="Y" in exI, simp)
apply (rule conjI, fast)
apply (rule_tac x="t" in exI)
apply (rule_tac x="sa" in exI)
apply (simp add: par_tr_sym)
apply (rule, simp add: in_failures)
apply (elim conjE exE)
apply (rule_tac x="Z" in exI)
apply (rule_tac x="Y" in exI, simp)
apply (rule conjI, fast)
apply (rule_tac x="t" in exI)
apply (rule_tac x="sa" in exI)
apply (simp add: par_tr_sym)
done
(*------------------*
| csp law |
*------------------*)
lemmas cspF_commut = cspF_Ext_choice_commut cspF_Int_choice_commut cspF_Parallel_commut
(*-----------------------------------*
| Associativity |
*-----------------------------------*)
lemma cspF_Ext_choice_assoc:
"P [+] (Q [+] R) =F[M,M] (P [+] Q) [+] R"
apply (simp add: cspF_cspT_semantics)
apply (simp add: cspT_Ext_choice_assoc)
apply (rule order_antisym)
apply (rule, simp add: in_failures in_traces)
apply (force)
apply (rule, simp add: in_failures in_traces)
apply (force)
done
lemma cspF_Ext_choice_assoc_sym:
"(P [+] Q) [+] R =F[M,M] P [+] (Q [+] R)"
apply (rule cspF_sym)
apply (simp add: cspF_Ext_choice_assoc)
done
lemma cspF_Int_choice_assoc:
"P |~| (Q |~| R) =F[M,M] (P |~| Q) |~| R"
apply (simp add: cspF_cspT_semantics)
apply (simp add: cspT_Int_choice_assoc)
apply (rule order_antisym)
apply (rule, simp add: in_failures)+
done
lemma cspF_Int_choice_assoc_sym:
"(P |~| Q) |~| R =F[M,M] P |~| (Q |~| R)"
apply (rule cspF_sym)
apply (simp add: cspF_Int_choice_assoc)
done
(*------------------*
| csp law |
*------------------*)
lemmas cspF_assoc = cspF_Ext_choice_assoc cspF_Int_choice_assoc
lemmas cspF_assoc_sym = cspF_Ext_choice_assoc_sym cspF_Int_choice_assoc_sym
(*-----------------------------------*
| Left Commutativity |
*-----------------------------------*)
lemma cspF_Ext_choice_left_commut:
"P [+] (Q [+] R) =F[M,M] Q [+] (P [+] R)"
apply (simp add: cspF_cspT_semantics)
apply (simp add: cspT_Ext_choice_left_commut)
apply (rule order_antisym)
apply (rule, simp add: in_failures in_traces)
apply (force)
apply (rule, simp add: in_failures in_traces)
apply (force)
done
lemma cspF_Int_choice_left_commut:
"P |~| (Q |~| R) =F[M,M] Q |~| (P |~| R)"
apply (simp add: cspF_cspT_semantics)
apply (simp add: cspT_Int_choice_left_commut)
apply (rule order_antisym)
apply (rule, simp add: in_failures)+
done
lemmas cspF_left_commut =
cspF_Ext_choice_left_commut cspF_Int_choice_left_commut
(*-----------------------------------*
| Unit |
*-----------------------------------*)
(*** STOP [+] P ***)
lemma cspF_Ext_choice_unit_l:
"STOP [+] P =F[M,M] P"
apply (simp add: cspF_cspT_semantics)
apply (simp add: cspT_Ext_choice_unit_l)
apply (rule order_antisym)
apply (rule, simp add: in_traces in_failures)
apply (elim conjE disjE)
apply (simp_all)
apply (rule proc_F2_F4)
apply (simp_all)
apply (rule, simp add: in_failures)
done
lemma cspF_Ext_choice_unit_r:
"P [+] STOP =F[M,M] P"
apply (rule cspF_rw_left)
apply (rule cspF_Ext_choice_commut)
apply (simp add: cspF_Ext_choice_unit_l)
done
lemmas cspF_Ext_choice_unit =
cspF_Ext_choice_unit_l cspF_Ext_choice_unit_r
lemma cspF_Int_choice_unit_l:
"DIV |~| P =F[M,M] P"
apply (simp add: cspF_cspT_semantics)
apply (simp add: cspT_Int_choice_unit_l)
apply (rule order_antisym)
apply (rule, simp add: in_failures)
apply (rule, simp add: in_failures)
done
lemma cspF_Int_choice_unit_r:
"P |~| DIV =F[M,M] P"
apply (rule cspF_rw_left)
apply (rule cspF_Int_choice_commut)
apply (simp add: cspF_Int_choice_unit_l)
done
lemmas cspF_Int_choice_unit =
cspF_Int_choice_unit_l cspF_Int_choice_unit_r
lemmas cspF_unit = cspF_Ext_choice_unit cspF_Int_choice_unit
(*-----------------------------------*
| !!-empty |
*-----------------------------------*)
lemma cspF_Rep_int_choice_sum_DIV:
"sumset C = {} ==> !! : C .. Pf =F[M1,M2] DIV"
apply (simp add: cspF_cspT_semantics)
apply (simp add: cspT_Rep_int_choice_DIV)
apply (simp add: failures_def)
apply (simp add: empF_def)
done
lemma cspF_Rep_int_choice_nat_DIV:
"!nat :{} .. Pf =F[M1,M2] DIV"
by (simp add: Rep_int_choice_ss_def cspF_Rep_int_choice_sum_DIV)
lemma cspF_Rep_int_choice_set_DIV:
"!set :{} .. Pf =F[M1,M2] DIV"
by (simp add: Rep_int_choice_ss_def cspF_Rep_int_choice_sum_DIV)
lemma cspF_Rep_int_choice_com_DIV:
"! :{} .. Pf =F[M1,M2] DIV"
by (simp add: Rep_int_choice_com_def cspF_Rep_int_choice_set_DIV)
lemma cspF_Rep_int_choice_f_DIV:
"inj f ==> !<f> :{} .. Pf =F[M1,M2] DIV"
by (simp add: Rep_int_choice_f_def cspF_Rep_int_choice_com_DIV)
lemmas cspF_Rep_int_choice_DIV = cspF_Rep_int_choice_sum_DIV
cspF_Rep_int_choice_nat_DIV
cspF_Rep_int_choice_set_DIV
cspF_Rep_int_choice_com_DIV
cspF_Rep_int_choice_f_DIV
lemmas cspF_Rep_int_choice_DIV_sym = cspF_Rep_int_choice_DIV[THEN cspF_sym]
lemmas cspF_Rep_int_choice_empty = cspF_Rep_int_choice_DIV
lemma cspF_DIV_top: "P <=F DIV"
apply (simp add: cspF_semantics)
apply (simp add: traces_def failures_def)
done
(*-----------------------------------*
| !!-unit |
*-----------------------------------*)
lemma cspF_Rep_int_choice_sum_unit:
"sumset C ~= {} ==> !! c:C .. P =F[M,M] P"
apply (simp add: cspF_cspT_semantics)
apply (simp add: cspT_Rep_int_choice_unit)
apply (rule order_antisym)
apply (rule, simp add: in_failures)
apply (rule, simp add: in_failures)
apply (force)
done
lemma cspF_Rep_int_choice_nat_unit:
"N ~= {} ==> !nat n:N .. P =F[M,M] P"
by (simp add: Rep_int_choice_ss_def cspF_Rep_int_choice_sum_unit)
lemma cspF_Rep_int_choice_set_unit:
"Xs ~= {} ==> !set X:Xs .. P =F[M,M] P"
by (simp add: Rep_int_choice_ss_def cspF_Rep_int_choice_sum_unit)
lemma cspF_Rep_int_choice_com_unit:
"X ~= {} ==> ! a:X .. P =F[M,M] P"
by (simp add: Rep_int_choice_com_def cspF_Rep_int_choice_set_unit)
lemma cspF_Rep_int_choice_f_unit:
"X ~= {} ==> !<f> a:X .. P =F[M,M] P"
by (simp add: Rep_int_choice_f_def cspF_Rep_int_choice_com_unit)
lemmas cspF_Rep_int_choice_unit =
cspF_Rep_int_choice_sum_unit
cspF_Rep_int_choice_nat_unit
cspF_Rep_int_choice_set_unit
cspF_Rep_int_choice_com_unit
cspF_Rep_int_choice_f_unit
(*-----------------------------------*
| !!-const |
*-----------------------------------*)
(* const *)
lemma cspF_Rep_int_choice_sum_const:
"[| sumset C ~= {} ; ALL c: sumset C. Pf c = P |] ==> !! :C .. Pf =F[M,M] P"
apply (simp add: cspF_cspT_semantics)
apply (simp add: cspT_Rep_int_choice_const)
apply (rule order_antisym)
apply (rule, simp add: in_failures)
apply (rule, simp add: in_failures)
apply (force)
done
lemma cspF_Rep_int_choice_nat_const:
"[| N ~= {} ; ALL n:N. Pf n = P |] ==> !nat :N .. Pf =F[M,M] P"
apply (simp add: Rep_int_choice_ss_def)
by (rule cspF_Rep_int_choice_sum_const, auto)
lemma cspF_Rep_int_choice_set_const:
"[| Xs ~= {} ; ALL X:Xs. Pf X = P |] ==> !set :Xs .. Pf =F[M,M] P"
apply (simp add: Rep_int_choice_ss_def)
by (rule cspF_Rep_int_choice_sum_const, auto)
lemma cspF_Rep_int_choice_com_const:
"[| X ~= {} ; ALL a:X. Pf a = P |] ==> ! :X .. Pf =F[M,M] P"
apply (simp add: Rep_int_choice_com_def)
by (rule cspF_Rep_int_choice_set_const, auto)
lemma cspF_Rep_int_choice_f_const:
"[| inj f ; X ~= {} ; ALL a:X. Pf a = P |] ==> !<f> :X .. Pf =F[M,M] P"
apply (simp add: Rep_int_choice_f_def)
by (rule cspF_Rep_int_choice_com_const, auto)
lemmas cspF_Rep_int_choice_const =
cspF_Rep_int_choice_sum_const
cspF_Rep_int_choice_nat_const
cspF_Rep_int_choice_set_const
cspF_Rep_int_choice_com_const
cspF_Rep_int_choice_f_const
(*-----------------------------------*
| |~|-!!-union |
*-----------------------------------*)
lemma cspF_Int_Rep_int_choice_sum_union:
"C1 =type= C2 ==>
(!! :C1 .. P1f) |~| (!! :C2 .. P2f)
=F[M,M] (!! c:(C1 Uns C2) ..
IF (c : sumset C1 & c : sumset C2) THEN (P1f c |~| P2f c)
ELSE IF (c : sumset C1) THEN P1f c ELSE P2f c)"
apply (simp add: cspF_cspT_semantics)
apply (simp add: cspT_Int_Rep_int_choice_union)
apply (rule order_antisym)
apply (rule)
apply (simp add: in_failures)
apply (elim conjE bexE disjE)
apply (rule_tac x="c" in bexI)
apply (simp)
apply (simp)
apply (rule_tac x="c" in bexI)
apply (simp)
apply (simp)
(* => *)
apply (rule)
apply (simp add: in_failures)
apply (elim conjE exE bexE)
apply (simp_all)
apply (elim disjE)
apply (simp_all)
apply (case_tac "c : sumset C2")
apply (simp add: in_failures)
apply (force)
apply (simp add: in_failures)
apply (force)
apply (case_tac "c : sumset C1")
apply (simp add: in_failures)
apply (force)
apply (simp add: in_failures)
apply (force)
done
lemma cspF_Int_Rep_int_choice_nat_union:
"(!nat :N1 .. P1f) |~| (!nat :N2 .. P2f)
=F[M,M] (!nat n:(N1 Un N2) ..
IF (n : N1 & n : N2) THEN (P1f n |~| P2f n)
ELSE IF (n : N1) THEN P1f n ELSE P2f n)"
apply (simp add: Rep_int_choice_ss_def)
apply (rule cspF_rw_left)
apply (rule cspF_Int_Rep_int_choice_sum_union)
apply (simp_all)
apply (rule cspF_decompo)
by (auto)
lemma cspF_Int_Rep_int_choice_set_union:
"(!set :Xs1 .. P1f) |~| (!set :Xs2 .. P2f)
=F[M,M] (!set X:(Xs1 Un Xs2) ..
IF (X : Xs1 & X : Xs2) THEN (P1f X |~| P2f X)
ELSE IF (X : Xs1) THEN P1f X ELSE P2f X)"
apply (simp add: Rep_int_choice_ss_def)
apply (rule cspF_rw_left)
apply (rule cspF_Int_Rep_int_choice_sum_union)
apply (simp_all)
apply (rule cspF_decompo)
by (auto)
lemma cspF_Int_Rep_int_choice_com_union:
"(! :X1 .. P1f) |~| (! :X2 .. P2f)
=F[M,M] (! a:(X1 Un X2) ..
IF (a : X1 & a : X2) THEN (P1f a |~| P2f a)
ELSE IF (a : X1) THEN P1f a ELSE P2f a)"
apply (simp add: Rep_int_choice_com_def)
apply (rule cspF_rw_left)
apply (rule cspF_Int_Rep_int_choice_set_union)
apply (simp_all)
apply (rule cspF_decompo)
by (auto)
lemma cspF_Int_Rep_int_choice_f_union:
"inj f ==>
(!<f> :X1 .. P1f) |~| (!<f> :X2 .. P2f)
=F[M,M] (!<f> a:(X1 Un X2) ..
IF (a : X1 & a : X2) THEN (P1f a |~| P2f a)
ELSE IF (a : X1) THEN P1f a ELSE P2f a)"
apply (simp add: Rep_int_choice_f_def)
apply (rule cspF_rw_left)
apply (rule cspF_Int_Rep_int_choice_com_union)
apply (rule cspF_decompo)
apply (auto simp add: inj_image_mem_iff)
done
lemmas cspF_Int_Rep_int_choice_union =
cspF_Int_Rep_int_choice_sum_union
cspF_Int_Rep_int_choice_nat_union
cspF_Int_Rep_int_choice_set_union
cspF_Int_Rep_int_choice_com_union
cspF_Int_Rep_int_choice_f_union
(*-----------------------------------*
| !!-union-|~| |
*-----------------------------------*)
lemma cspF_Rep_int_choice_sum_union_Int:
"C1 =type= C2 ==>
(!! :(C1 Uns C2) .. Pf)
=F[M,M] (!! c:C1 .. Pf c) |~| (!! c:C2 .. Pf c)"
apply (rule cspF_rw_right)
apply (rule cspF_Int_Rep_int_choice_union)
apply (simp)
apply (rule cspF_decompo)
apply (simp)
apply (rule cspF_rw_right)
apply (rule cspF_IF_split)
apply (simp)
apply (simp add: cspF_idem[THEN cspF_sym])
apply (intro impI)
apply (rule cspF_rw_right)
apply (rule cspF_IF_split)
apply (simp)
done
lemma cspF_Rep_int_choice_nat_union_Int:
"(!nat :(N1 Un N2) .. Pf)
=F[M,M] (!nat n:N1 .. Pf n) |~| (!nat n:N2 .. Pf n)"
apply (simp add: Rep_int_choice_ss_def)
apply (rule cspF_rw_right)
apply (rule cspF_Rep_int_choice_sum_union_Int[THEN cspF_sym])
apply (simp_all)
done
lemma cspF_Rep_int_choice_set_union_Int:
"(!set :(Xs1 Un Xs2) .. Pf)
=F[M,M] (!set X:Xs1 .. Pf X) |~| (!set X:Xs2 .. Pf X)"
apply (simp add: Rep_int_choice_ss_def)
apply (rule cspF_rw_right)
apply (rule cspF_Rep_int_choice_sum_union_Int[THEN cspF_sym])
apply (simp_all)
done
lemma cspF_Rep_int_choice_com_union_Int:
"(! :(X1 Un X2) .. Pf)
=F[M,M] (! a:X1 .. Pf a) |~| (! a:X2 .. Pf a)"
apply (simp add: Rep_int_choice_com_def)
apply (rule cspF_rw_right)
apply (rule cspF_Rep_int_choice_set_union_Int[THEN cspF_sym])
apply (rule cspF_decompo)
apply (auto)
done
lemma cspF_Rep_int_choice_f_union_Int:
"(!<f> :(X1 Un X2) .. Pf)
=F[M,M] (!<f> x:X1 .. Pf x) |~| (!<f> x:X2 .. Pf x)"
apply (simp add: Rep_int_choice_f_def)
apply (rule cspF_rw_right)
apply (rule cspF_Rep_int_choice_com_union_Int[THEN cspF_sym])
apply (rule cspF_decompo)
apply (auto)
done
lemmas cspF_Rep_int_choice_union_Int =
cspF_Rep_int_choice_sum_union_Int
cspF_Rep_int_choice_nat_union_Int
cspF_Rep_int_choice_set_union_Int
cspF_Rep_int_choice_com_union_Int
cspF_Rep_int_choice_f_union_Int
(*********************************************************
Depth_rest
*********************************************************)
(*------------------*
| csp law |
*------------------*)
lemma cspF_Depth_rest_Zero:
"P |. 0 =F[M1,M2] DIV"
apply (simp add: cspF_cspT_semantics)
apply (simp add: cspT_Depth_rest_Zero)
apply (rule order_antisym)
(* => *)
apply (rule)
apply (simp add: in_failures)
apply (force)
(* <= *)
apply (rule)
apply (simp add: in_failures)
done
lemma cspF_Depth_rest_min:
"P |. n |. m =F[M,M] P |. min n m"
apply (simp add: cspF_cspT_semantics)
apply (simp add: cspT_Depth_rest_min)
apply (simp add: failures.simps)
apply (simp add: min_rs)
done
lemma cspF_Depth_rest_congE:
"[| P =F[M1,M2] Q ; ALL m. P |. m =F[M1,M2] Q |. m ==> S |] ==> S"
apply (simp add: cspF_semantics)
apply (simp add: traces.simps)
apply (simp add: failures.simps)
done
lemma cspF_Depth_rest_n:
"P |. n |. n =F[M,M] P |. n"
apply (rule cspF_rw_left)
apply (rule cspF_Depth_rest_min)
apply (simp)
done
(*------------------*
| !nat-rest |
*------------------*)
lemma cspF_nat_Depth_rest_UNIV:
"P =F[M,M] !nat n .. (P |. n)"
apply (simp add: cspF_cspT_semantics)
apply (simp add: cspT_nat_Depth_rest_UNIV)
apply (rule order_antisym)
(* <= *)
apply (rule)
apply (simp add: in_failures)
apply (case_tac "noTick s")
apply (rule_tac x="Suc (lengtht s)" in exI)
apply (simp)
apply (rule_tac x="lengtht s" in exI)
apply (simp)
apply (rule_tac x="(butlastt s)" in exI)
apply (simp add: Tick_decompo)
apply (simp add: noTick_butlast)
(* => *)
apply (rule)
apply (simp add: in_failures)
done
lemma cspF_nat_Depth_rest_lengthset:
"P =F[M,M] !nat n:(lengthset P (fstF o M)) .. (P |. n)"
apply (simp add: cspF_cspT_semantics)
apply (simp add: cspT_nat_Depth_rest_lengthset)
apply (rule order_antisym)
(* <= *)
apply (rule)
apply (simp add: in_failures)
apply (case_tac "noTick s")
apply (rule_tac x="Suc (lengtht s)" in bexI)
apply (simp)
apply (simp add: lengthset_def)
apply (rule_tac x="s" in exI)
apply (simp add: proc_T2)
apply (rule_tac x="lengtht s" in bexI)
apply (simp)
apply (rule_tac x="(butlastt s)" in exI)
apply (simp add: Tick_decompo)
apply (simp add: noTick_butlast)
apply (simp add: lengthset_def)
apply (rule_tac x="s" in exI)
apply (simp add: proc_T2)
(* => *)
apply (rule)
apply (simp add: in_failures)
done
lemmas cspF_nat_Depth_rest = cspF_nat_Depth_rest_UNIV
cspF_nat_Depth_rest_lengthset
(*------------------*
| ?-partial |
*------------------*)
lemma cspF_Ext_pre_choice_partial:
"? :X -> Pf =F[M,M] ? x:X -> (IF (x:X) THEN Pf x ELSE DIV)"
apply (rule cspF_decompo)
apply (simp_all)
apply (rule cspF_rw_right)
apply (rule cspF_IF)
apply (simp)
done
(*------------------*
| !!-partial |
*------------------*)
lemma cspF_Rep_int_choice_sum_partial:
"!! :C .. Pf =F[M,M] !! c:C .. (IF (c: sumset C) THEN Pf c ELSE DIV)"
apply (rule cspF_decompo)
apply (simp_all)
apply (rule cspF_rw_right)
apply (rule cspF_IF)
apply (simp)
done
lemma cspF_Rep_int_choice_nat_partial:
"!nat :N .. Pf =F[M,M] !nat n:N .. (IF (n:N) THEN Pf n ELSE DIV)"
apply (rule cspF_decompo)
apply (simp_all)
apply (rule cspF_rw_right)
apply (rule cspF_IF)
apply (simp)
done
lemma cspF_Rep_int_choice_set_partial:
"!set :Xs .. Pf =F[M,M] !set X:Xs .. (IF (X:Xs) THEN Pf X ELSE DIV)"
apply (rule cspF_decompo)
apply (simp_all)
apply (rule cspF_rw_right)
apply (rule cspF_IF)
apply (simp)
done
lemma cspF_Rep_int_choice_com_partial:
"! :X .. Pf =F[M,M] ! a:X .. (IF (a:X) THEN Pf a ELSE DIV)"
apply (rule cspF_decompo)
apply (simp_all)
apply (rule cspF_rw_right)
apply (rule cspF_IF)
apply (simp)
done
lemma cspF_Rep_int_choice_f_partial:
"inj f ==> !<f> :X .. Pf =F[M,M] !<f> a:X .. (IF (a:X) THEN Pf a ELSE DIV)"
apply (rule cspF_decompo)
apply (simp_all)
apply (rule cspF_rw_right)
apply (rule cspF_IF)
apply (simp)
done
lemmas cspF_Rep_int_choice_partial =
cspF_Rep_int_choice_sum_partial
cspF_Rep_int_choice_nat_partial
cspF_Rep_int_choice_set_partial
cspF_Rep_int_choice_com_partial
cspF_Rep_int_choice_f_partial
end
lemma cspF_IF_split:
IF b THEN P ELSE Q =F[M,M] (if b then P else Q)
lemma cspF_IF_True:
IF True THEN P ELSE Q =F[M,M] P
lemma cspF_IF_False:
IF False THEN P ELSE Q =F[M,M] Q
lemmas cspF_IF:
IF True THEN P ELSE Q =F[M,M] P
IF False THEN P ELSE Q =F[M,M] Q
lemmas cspF_IF:
IF True THEN P ELSE Q =F[M,M] P
IF False THEN P ELSE Q =F[M,M] Q
lemma cspF_Ext_choice_idem:
P [+] P =F[M,M] P
lemma cspF_Int_choice_idem:
P |~| P =F[M,M] P
lemmas cspF_idem:
P [+] P =F[M,M] P
P |~| P =F[M,M] P
lemmas cspF_idem:
P [+] P =F[M,M] P
P |~| P =F[M,M] P
lemma cspF_Ext_choice_commut:
P [+] Q =F[M,M] Q [+] P
lemma cspF_Int_choice_commut:
P |~| Q =F[M,M] Q |~| P
lemma cspF_Parallel_commut:
P |[X]| Q =F[M,M] Q |[X]| P
lemmas cspF_commut:
P [+] Q =F[M,M] Q [+] P
P |~| Q =F[M,M] Q |~| P
P |[X]| Q =F[M,M] Q |[X]| P
lemmas cspF_commut:
P [+] Q =F[M,M] Q [+] P
P |~| Q =F[M,M] Q |~| P
P |[X]| Q =F[M,M] Q |[X]| P
lemma cspF_Ext_choice_assoc:
P [+] (Q [+] R) =F[M,M] P [+] Q [+] R
lemma cspF_Ext_choice_assoc_sym:
P [+] Q [+] R =F[M,M] P [+] (Q [+] R)
lemma cspF_Int_choice_assoc:
P |~| (Q |~| R) =F[M,M] P |~| Q |~| R
lemma cspF_Int_choice_assoc_sym:
P |~| Q |~| R =F[M,M] P |~| (Q |~| R)
lemmas cspF_assoc:
P [+] (Q [+] R) =F[M,M] P [+] Q [+] R
P |~| (Q |~| R) =F[M,M] P |~| Q |~| R
lemmas cspF_assoc:
P [+] (Q [+] R) =F[M,M] P [+] Q [+] R
P |~| (Q |~| R) =F[M,M] P |~| Q |~| R
lemmas cspF_assoc_sym:
P [+] Q [+] R =F[M,M] P [+] (Q [+] R)
P |~| Q |~| R =F[M,M] P |~| (Q |~| R)
lemmas cspF_assoc_sym:
P [+] Q [+] R =F[M,M] P [+] (Q [+] R)
P |~| Q |~| R =F[M,M] P |~| (Q |~| R)
lemma cspF_Ext_choice_left_commut:
P [+] (Q [+] R) =F[M,M] Q [+] (P [+] R)
lemma cspF_Int_choice_left_commut:
P |~| (Q |~| R) =F[M,M] Q |~| (P |~| R)
lemmas cspF_left_commut:
P [+] (Q [+] R) =F[M,M] Q [+] (P [+] R)
P |~| (Q |~| R) =F[M,M] Q |~| (P |~| R)
lemmas cspF_left_commut:
P [+] (Q [+] R) =F[M,M] Q [+] (P [+] R)
P |~| (Q |~| R) =F[M,M] Q |~| (P |~| R)
lemma cspF_Ext_choice_unit_l:
STOP [+] P =F[M,M] P
lemma cspF_Ext_choice_unit_r:
P [+] STOP =F[M,M] P
lemmas cspF_Ext_choice_unit:
STOP [+] P =F[M,M] P
P [+] STOP =F[M,M] P
lemmas cspF_Ext_choice_unit:
STOP [+] P =F[M,M] P
P [+] STOP =F[M,M] P
lemma cspF_Int_choice_unit_l:
DIV |~| P =F[M,M] P
lemma cspF_Int_choice_unit_r:
P |~| DIV =F[M,M] P
lemmas cspF_Int_choice_unit:
DIV |~| P =F[M,M] P
P |~| DIV =F[M,M] P
lemmas cspF_Int_choice_unit:
DIV |~| P =F[M,M] P
P |~| DIV =F[M,M] P
lemmas cspF_unit:
STOP [+] P =F[M,M] P
P [+] STOP =F[M,M] P
DIV |~| P =F[M,M] P
P |~| DIV =F[M,M] P
lemmas cspF_unit:
STOP [+] P =F[M,M] P
P [+] STOP =F[M,M] P
DIV |~| P =F[M,M] P
P |~| DIV =F[M,M] P
lemma cspF_Rep_int_choice_sum_DIV:
sumset C = {} ==> !! :C .. Pf =F[M1.0,M2.0] DIV
lemma cspF_Rep_int_choice_nat_DIV:
!nat :{} .. Pf =F[M1.0,M2.0] DIV
lemma cspF_Rep_int_choice_set_DIV:
!set :{} .. Pf =F[M1.0,M2.0] DIV
lemma cspF_Rep_int_choice_com_DIV:
! :{} .. Pf =F[M1.0,M2.0] DIV
lemma cspF_Rep_int_choice_f_DIV:
inj f ==> !<f> :{} .. Pf =F[M1.0,M2.0] DIV
lemmas cspF_Rep_int_choice_DIV:
sumset C = {} ==> !! :C .. Pf =F[M1.0,M2.0] DIV
!nat :{} .. Pf =F[M1.0,M2.0] DIV
!set :{} .. Pf =F[M1.0,M2.0] DIV
! :{} .. Pf =F[M1.0,M2.0] DIV
inj f ==> !<f> :{} .. Pf =F[M1.0,M2.0] DIV
lemmas cspF_Rep_int_choice_DIV:
sumset C = {} ==> !! :C .. Pf =F[M1.0,M2.0] DIV
!nat :{} .. Pf =F[M1.0,M2.0] DIV
!set :{} .. Pf =F[M1.0,M2.0] DIV
! :{} .. Pf =F[M1.0,M2.0] DIV
inj f ==> !<f> :{} .. Pf =F[M1.0,M2.0] DIV
lemmas cspF_Rep_int_choice_DIV_sym:
sumset C1 = {} ==> DIV =F[M2.0,M1.0] !! :C1 .. Pf1
DIV =F[M2.0,M1.0] !nat :{} .. Pf1
DIV =F[M2.0,M1.0] !set :{} .. Pf1
DIV =F[M2.0,M1.0] ! :{} .. Pf1
inj f1 ==> DIV =F[M2.0,M1.0] !<f1> :{} .. Pf1
lemmas cspF_Rep_int_choice_DIV_sym:
sumset C1 = {} ==> DIV =F[M2.0,M1.0] !! :C1 .. Pf1
DIV =F[M2.0,M1.0] !nat :{} .. Pf1
DIV =F[M2.0,M1.0] !set :{} .. Pf1
DIV =F[M2.0,M1.0] ! :{} .. Pf1
inj f1 ==> DIV =F[M2.0,M1.0] !<f1> :{} .. Pf1
lemmas cspF_Rep_int_choice_empty:
sumset C = {} ==> !! :C .. Pf =F[M1.0,M2.0] DIV
!nat :{} .. Pf =F[M1.0,M2.0] DIV
!set :{} .. Pf =F[M1.0,M2.0] DIV
! :{} .. Pf =F[M1.0,M2.0] DIV
inj f ==> !<f> :{} .. Pf =F[M1.0,M2.0] DIV
lemmas cspF_Rep_int_choice_empty:
sumset C = {} ==> !! :C .. Pf =F[M1.0,M2.0] DIV
!nat :{} .. Pf =F[M1.0,M2.0] DIV
!set :{} .. Pf =F[M1.0,M2.0] DIV
! :{} .. Pf =F[M1.0,M2.0] DIV
inj f ==> !<f> :{} .. Pf =F[M1.0,M2.0] DIV
lemma cspF_DIV_top:
P <=F DIV
lemma cspF_Rep_int_choice_sum_unit:
sumset C ≠ {} ==> !! c:C .. P =F[M,M] P
lemma cspF_Rep_int_choice_nat_unit:
N ≠ {} ==> !nat n:N .. P =F[M,M] P
lemma cspF_Rep_int_choice_set_unit:
Xs ≠ {} ==> !set X:Xs .. P =F[M,M] P
lemma cspF_Rep_int_choice_com_unit:
X ≠ {} ==> ! a:X .. P =F[M,M] P
lemma cspF_Rep_int_choice_f_unit:
X ≠ {} ==> !<f> a:X .. P =F[M,M] P
lemmas cspF_Rep_int_choice_unit:
sumset C ≠ {} ==> !! c:C .. P =F[M,M] P
N ≠ {} ==> !nat n:N .. P =F[M,M] P
Xs ≠ {} ==> !set X:Xs .. P =F[M,M] P
X ≠ {} ==> ! a:X .. P =F[M,M] P
X ≠ {} ==> !<f> a:X .. P =F[M,M] P
lemmas cspF_Rep_int_choice_unit:
sumset C ≠ {} ==> !! c:C .. P =F[M,M] P
N ≠ {} ==> !nat n:N .. P =F[M,M] P
Xs ≠ {} ==> !set X:Xs .. P =F[M,M] P
X ≠ {} ==> ! a:X .. P =F[M,M] P
X ≠ {} ==> !<f> a:X .. P =F[M,M] P
lemma cspF_Rep_int_choice_sum_const:
[| sumset C ≠ {}; ∀c∈sumset C. Pf c = P |] ==> !! :C .. Pf =F[M,M] P
lemma cspF_Rep_int_choice_nat_const:
[| N ≠ {}; ∀n∈N. Pf n = P |] ==> !nat :N .. Pf =F[M,M] P
lemma cspF_Rep_int_choice_set_const:
[| Xs ≠ {}; ∀X∈Xs. Pf X = P |] ==> !set :Xs .. Pf =F[M,M] P
lemma cspF_Rep_int_choice_com_const:
[| X ≠ {}; ∀a∈X. Pf a = P |] ==> ! :X .. Pf =F[M,M] P
lemma cspF_Rep_int_choice_f_const:
[| inj f; X ≠ {}; ∀a∈X. Pf a = P |] ==> !<f> :X .. Pf =F[M,M] P
lemmas cspF_Rep_int_choice_const:
[| sumset C ≠ {}; ∀c∈sumset C. Pf c = P |] ==> !! :C .. Pf =F[M,M] P
[| N ≠ {}; ∀n∈N. Pf n = P |] ==> !nat :N .. Pf =F[M,M] P
[| Xs ≠ {}; ∀X∈Xs. Pf X = P |] ==> !set :Xs .. Pf =F[M,M] P
[| X ≠ {}; ∀a∈X. Pf a = P |] ==> ! :X .. Pf =F[M,M] P
[| inj f; X ≠ {}; ∀a∈X. Pf a = P |] ==> !<f> :X .. Pf =F[M,M] P
lemmas cspF_Rep_int_choice_const:
[| sumset C ≠ {}; ∀c∈sumset C. Pf c = P |] ==> !! :C .. Pf =F[M,M] P
[| N ≠ {}; ∀n∈N. Pf n = P |] ==> !nat :N .. Pf =F[M,M] P
[| Xs ≠ {}; ∀X∈Xs. Pf X = P |] ==> !set :Xs .. Pf =F[M,M] P
[| X ≠ {}; ∀a∈X. Pf a = P |] ==> ! :X .. Pf =F[M,M] P
[| inj f; X ≠ {}; ∀a∈X. Pf a = P |] ==> !<f> :X .. Pf =F[M,M] P
lemma cspF_Int_Rep_int_choice_sum_union:
C1.0 =type= C2.0 ==> !! :C1.0 .. P1f |~| !! :C2.0 .. P2f =F[M,M] !! c:(C1.0 Uns C2.0) .. IF (c ∈ sumset C1.0 ∧ c ∈ sumset C2.0) THEN P1f c |~| P2f c ELSE IF (c ∈ sumset C1.0) THEN P1f c ELSE P2f c
lemma cspF_Int_Rep_int_choice_nat_union:
!nat :N1.0 .. P1f |~| !nat :N2.0 .. P2f =F[M,M] !nat n:(N1.0 ∪ N2.0) .. IF (n ∈ N1.0 ∧ n ∈ N2.0) THEN P1f n |~| P2f n ELSE IF (n ∈ N1.0) THEN P1f n ELSE P2f n
lemma cspF_Int_Rep_int_choice_set_union:
!set :Xs1.0 .. P1f |~| !set :Xs2.0 .. P2f =F[M,M] !set X:(Xs1.0 ∪ Xs2.0) .. IF (X ∈ Xs1.0 ∧ X ∈ Xs2.0) THEN P1f X |~| P2f X ELSE IF (X ∈ Xs1.0) THEN P1f X ELSE P2f X
lemma cspF_Int_Rep_int_choice_com_union:
! :X1.0 .. P1f |~| ! :X2.0 .. P2f =F[M,M] ! a:(X1.0 ∪ X2.0) .. IF (a ∈ X1.0 ∧ a ∈ X2.0) THEN P1f a |~| P2f a ELSE IF (a ∈ X1.0) THEN P1f a ELSE P2f a
lemma cspF_Int_Rep_int_choice_f_union:
inj f ==> !<f> :X1.0 .. P1f |~| !<f> :X2.0 .. P2f =F[M,M] !<f> a:(X1.0 ∪ X2.0) .. IF (a ∈ X1.0 ∧ a ∈ X2.0) THEN P1f a |~| P2f a ELSE IF (a ∈ X1.0) THEN P1f a ELSE P2f a
lemmas cspF_Int_Rep_int_choice_union:
C1.0 =type= C2.0 ==> !! :C1.0 .. P1f |~| !! :C2.0 .. P2f =F[M,M] !! c:(C1.0 Uns C2.0) .. IF (c ∈ sumset C1.0 ∧ c ∈ sumset C2.0) THEN P1f c |~| P2f c ELSE IF (c ∈ sumset C1.0) THEN P1f c ELSE P2f c
!nat :N1.0 .. P1f |~| !nat :N2.0 .. P2f =F[M,M] !nat n:(N1.0 ∪ N2.0) .. IF (n ∈ N1.0 ∧ n ∈ N2.0) THEN P1f n |~| P2f n ELSE IF (n ∈ N1.0) THEN P1f n ELSE P2f n
!set :Xs1.0 .. P1f |~| !set :Xs2.0 .. P2f =F[M,M] !set X:(Xs1.0 ∪ Xs2.0) .. IF (X ∈ Xs1.0 ∧ X ∈ Xs2.0) THEN P1f X |~| P2f X ELSE IF (X ∈ Xs1.0) THEN P1f X ELSE P2f X
! :X1.0 .. P1f |~| ! :X2.0 .. P2f =F[M,M] ! a:(X1.0 ∪ X2.0) .. IF (a ∈ X1.0 ∧ a ∈ X2.0) THEN P1f a |~| P2f a ELSE IF (a ∈ X1.0) THEN P1f a ELSE P2f a
inj f ==> !<f> :X1.0 .. P1f |~| !<f> :X2.0 .. P2f =F[M,M] !<f> a:(X1.0 ∪ X2.0) .. IF (a ∈ X1.0 ∧ a ∈ X2.0) THEN P1f a |~| P2f a ELSE IF (a ∈ X1.0) THEN P1f a ELSE P2f a
lemmas cspF_Int_Rep_int_choice_union:
C1.0 =type= C2.0 ==> !! :C1.0 .. P1f |~| !! :C2.0 .. P2f =F[M,M] !! c:(C1.0 Uns C2.0) .. IF (c ∈ sumset C1.0 ∧ c ∈ sumset C2.0) THEN P1f c |~| P2f c ELSE IF (c ∈ sumset C1.0) THEN P1f c ELSE P2f c
!nat :N1.0 .. P1f |~| !nat :N2.0 .. P2f =F[M,M] !nat n:(N1.0 ∪ N2.0) .. IF (n ∈ N1.0 ∧ n ∈ N2.0) THEN P1f n |~| P2f n ELSE IF (n ∈ N1.0) THEN P1f n ELSE P2f n
!set :Xs1.0 .. P1f |~| !set :Xs2.0 .. P2f =F[M,M] !set X:(Xs1.0 ∪ Xs2.0) .. IF (X ∈ Xs1.0 ∧ X ∈ Xs2.0) THEN P1f X |~| P2f X ELSE IF (X ∈ Xs1.0) THEN P1f X ELSE P2f X
! :X1.0 .. P1f |~| ! :X2.0 .. P2f =F[M,M] ! a:(X1.0 ∪ X2.0) .. IF (a ∈ X1.0 ∧ a ∈ X2.0) THEN P1f a |~| P2f a ELSE IF (a ∈ X1.0) THEN P1f a ELSE P2f a
inj f ==> !<f> :X1.0 .. P1f |~| !<f> :X2.0 .. P2f =F[M,M] !<f> a:(X1.0 ∪ X2.0) .. IF (a ∈ X1.0 ∧ a ∈ X2.0) THEN P1f a |~| P2f a ELSE IF (a ∈ X1.0) THEN P1f a ELSE P2f a
lemma cspF_Rep_int_choice_sum_union_Int:
C1.0 =type= C2.0 ==> !! :(C1.0 Uns C2.0) .. Pf =F[M,M] !! :C1.0 .. Pf |~| !! :C2.0 .. Pf
lemma cspF_Rep_int_choice_nat_union_Int:
!nat :(N1.0 ∪ N2.0) .. Pf =F[M,M] !nat :N1.0 .. Pf |~| !nat :N2.0 .. Pf
lemma cspF_Rep_int_choice_set_union_Int:
!set :(Xs1.0 ∪ Xs2.0) .. Pf =F[M,M] !set :Xs1.0 .. Pf |~| !set :Xs2.0 .. Pf
lemma cspF_Rep_int_choice_com_union_Int:
! :(X1.0 ∪ X2.0) .. Pf =F[M,M] ! :X1.0 .. Pf |~| ! :X2.0 .. Pf
lemma cspF_Rep_int_choice_f_union_Int:
!<f> :(X1.0 ∪ X2.0) .. Pf =F[M,M] !<f> :X1.0 .. Pf |~| !<f> :X2.0 .. Pf
lemmas cspF_Rep_int_choice_union_Int:
C1.0 =type= C2.0 ==> !! :(C1.0 Uns C2.0) .. Pf =F[M,M] !! :C1.0 .. Pf |~| !! :C2.0 .. Pf
!nat :(N1.0 ∪ N2.0) .. Pf =F[M,M] !nat :N1.0 .. Pf |~| !nat :N2.0 .. Pf
!set :(Xs1.0 ∪ Xs2.0) .. Pf =F[M,M] !set :Xs1.0 .. Pf |~| !set :Xs2.0 .. Pf
! :(X1.0 ∪ X2.0) .. Pf =F[M,M] ! :X1.0 .. Pf |~| ! :X2.0 .. Pf
!<f> :(X1.0 ∪ X2.0) .. Pf =F[M,M] !<f> :X1.0 .. Pf |~| !<f> :X2.0 .. Pf
lemmas cspF_Rep_int_choice_union_Int:
C1.0 =type= C2.0 ==> !! :(C1.0 Uns C2.0) .. Pf =F[M,M] !! :C1.0 .. Pf |~| !! :C2.0 .. Pf
!nat :(N1.0 ∪ N2.0) .. Pf =F[M,M] !nat :N1.0 .. Pf |~| !nat :N2.0 .. Pf
!set :(Xs1.0 ∪ Xs2.0) .. Pf =F[M,M] !set :Xs1.0 .. Pf |~| !set :Xs2.0 .. Pf
! :(X1.0 ∪ X2.0) .. Pf =F[M,M] ! :X1.0 .. Pf |~| ! :X2.0 .. Pf
!<f> :(X1.0 ∪ X2.0) .. Pf =F[M,M] !<f> :X1.0 .. Pf |~| !<f> :X2.0 .. Pf
lemma cspF_Depth_rest_Zero:
P |. 0 =F[M1.0,M2.0] DIV
lemma cspF_Depth_rest_min:
P |. n |. m =F[M,M] P |. min n m
lemma cspF_Depth_rest_congE:
[| P =F[M1.0,M2.0] Q; ∀m. P |. m =F[M1.0,M2.0] Q |. m ==> S |] ==> S
lemma cspF_Depth_rest_n:
P |. n |. n =F[M,M] P |. n
lemma cspF_nat_Depth_rest_UNIV:
P =F[M,M] !nat :UNIV .. Depth_rest P
lemma cspF_nat_Depth_rest_lengthset:
P =F[M,M] !nat :lengthset P (fstF o M) .. Depth_rest P
lemmas cspF_nat_Depth_rest:
P =F[M,M] !nat :UNIV .. Depth_rest P
P =F[M,M] !nat :lengthset P (fstF o M) .. Depth_rest P
lemmas cspF_nat_Depth_rest:
P =F[M,M] !nat :UNIV .. Depth_rest P
P =F[M,M] !nat :lengthset P (fstF o M) .. Depth_rest P
lemma cspF_Ext_pre_choice_partial:
? :X -> Pf =F[M,M] ? x:X -> IF (x ∈ X) THEN Pf x ELSE DIV
lemma cspF_Rep_int_choice_sum_partial:
!! :C .. Pf =F[M,M] !! c:C .. IF (c ∈ sumset C) THEN Pf c ELSE DIV
lemma cspF_Rep_int_choice_nat_partial:
!nat :N .. Pf =F[M,M] !nat n:N .. IF (n ∈ N) THEN Pf n ELSE DIV
lemma cspF_Rep_int_choice_set_partial:
!set :Xs .. Pf =F[M,M] !set X:Xs .. IF (X ∈ Xs) THEN Pf X ELSE DIV
lemma cspF_Rep_int_choice_com_partial:
! :X .. Pf =F[M,M] ! a:X .. IF (a ∈ X) THEN Pf a ELSE DIV
lemma cspF_Rep_int_choice_f_partial:
inj f ==> !<f> :X .. Pf =F[M,M] !<f> a:X .. IF (a ∈ X) THEN Pf a ELSE DIV
lemmas cspF_Rep_int_choice_partial:
!! :C .. Pf =F[M,M] !! c:C .. IF (c ∈ sumset C) THEN Pf c ELSE DIV
!nat :N .. Pf =F[M,M] !nat n:N .. IF (n ∈ N) THEN Pf n ELSE DIV
!set :Xs .. Pf =F[M,M] !set X:Xs .. IF (X ∈ Xs) THEN Pf X ELSE DIV
! :X .. Pf =F[M,M] ! a:X .. IF (a ∈ X) THEN Pf a ELSE DIV
inj f ==> !<f> :X .. Pf =F[M,M] !<f> a:X .. IF (a ∈ X) THEN Pf a ELSE DIV
lemmas cspF_Rep_int_choice_partial:
!! :C .. Pf =F[M,M] !! c:C .. IF (c ∈ sumset C) THEN Pf c ELSE DIV
!nat :N .. Pf =F[M,M] !nat n:N .. IF (n ∈ N) THEN Pf n ELSE DIV
!set :Xs .. Pf =F[M,M] !set X:Xs .. IF (X ∈ Xs) THEN Pf X ELSE DIV
! :X .. Pf =F[M,M] ! a:X .. IF (a ∈ X) THEN Pf a ELSE DIV
inj f ==> !<f> :X .. Pf =F[M,M] !<f> a:X .. IF (a ∈ X) THEN Pf a ELSE DIV