New_obj_prop.thy
author chunhan
Thu, 09 Jan 2014 14:39:00 +0800
changeset 92 d9dc04c3ea90
parent 1 7d9c0ed02b56
permissions -rw-r--r--
modify co2sobj/s2ss from object to dobject
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     1
theory New_obj_prop
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     2
imports Main Finite_current Flask_type Flask Static
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     3
begin
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     4
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     5
context tainting_s begin
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     6
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     7
lemma nn_notin_aux: "finite s \<Longrightarrow> \<forall> a \<in> s. Max s \<ge> a "
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     8
apply (erule finite.induct, simp)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     9
apply (rule ballI)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    10
apply (case_tac "aa = a", simp+)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    11
done
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    12
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    13
lemma nn_notin: "finite s \<Longrightarrow> next_nat s \<notin> s"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    14
apply (drule nn_notin_aux)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    15
apply (simp add:next_nat_def)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    16
by (auto)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    17
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    18
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    19
(* lemmas of new created obj *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    20
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    21
lemma np_notin_curp: "valid \<tau> \<Longrightarrow> new_proc \<tau> \<notin> current_procs \<tau>" using finite_cp
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    22
by (simp add:new_proc_def nn_notin)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    23
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    24
lemma np_notin_curp': "\<lbrakk>new_proc \<tau> \<in> current_procs \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> False"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    25
apply (drule np_notin_curp, simp)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    26
done
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    27
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    28
lemma ni_notin_curi: "valid \<tau> \<Longrightarrow> new_inode_num \<tau> \<notin> current_inode_nums \<tau>"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    29
apply (drule finite_ci)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    30
by (simp add:new_inode_num_def nn_notin)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    31
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    32
lemma ni_notin_curi': "\<lbrakk>new_inode_num \<tau> \<in> current_inode_nums \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> False"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    33
by (drule ni_notin_curi, simp)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    34
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    35
lemma nm_notin_curm: "valid \<tau> \<Longrightarrow> new_msgq \<tau> \<notin> current_msgqs \<tau>" using finite_cm
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    36
by (simp add:new_msgq_def nn_notin)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    37
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    38
lemma nm_notin_curm': "\<lbrakk>new_msgq \<tau> \<in> current_msgqs \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> False"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    39
by (drule nm_notin_curm, simp)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    40
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    41
lemma nh_notin_curh: "valid \<tau> \<Longrightarrow> new_shm \<tau> \<notin> current_shms \<tau>" using finite_ch
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    42
by (simp add:new_shm_def nn_notin)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    43
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    44
lemma nh_notin_curh': "\<lbrakk>new_shm \<tau> \<in> current_shms \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> False"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    45
by (drule nh_notin_curh, simp)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    46
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    47
lemma nfd_notin_curfd: "valid \<tau> \<Longrightarrow> new_proc_fd \<tau> p \<notin> current_proc_fds \<tau> p" 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    48
using finite_cfd[where p = p]
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    49
apply (simp add:new_proc_fd_def nn_notin)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    50
done
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    51
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    52
lemma nfd_notin_curfd': "\<lbrakk>new_proc_fd \<tau> p \<in> current_proc_fds \<tau> p; valid \<tau>\<rbrakk> \<Longrightarrow> False"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    53
by (drule nfd_notin_curfd[where p = p], simp)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    54
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    55
lemma nim_notin_curim: "valid \<tau> \<Longrightarrow> new_inode_num \<tau> \<notin> current_inode_nums \<tau>"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    56
by (drule finite_ci, simp add:new_inode_num_def nn_notin)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    57
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    58
lemma nim_notin_curim': "\<lbrakk>new_inode_num \<tau> \<in> current_inode_nums \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> False"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    59
by (drule nim_notin_curim, simp)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    60
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    61
lemma len_fname_all: "length (fname_all_a len) = len"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    62
by (induct len, auto simp:fname_all_a.simps)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    63
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    64
lemma ncf_notin_curf: "valid \<tau> \<Longrightarrow> new_childf f \<tau> \<notin> current_files \<tau>"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    65
apply (drule finite_cf)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    66
apply (simp add:new_childf_def next_fname_def all_fname_under_dir_def)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    67
apply (rule notI)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    68
apply (subgoal_tac "(CHR ''a'' # fname_all_a (Max (fname_length_set {fn. fn # f \<in> current_files \<tau>}))) \<in> {fn. fn # f \<in> current_files \<tau>}")
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    69
defer apply simp
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    70
apply (subgoal_tac "length (CHR ''a'' # fname_all_a (Max (fname_length_set {fn. fn # f \<in> current_files \<tau>}))) \<in> fname_length_set {fn. fn # f \<in> current_files \<tau>}")
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    71
defer apply (auto simp:fname_length_set_def image_def)[1]
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    72
apply (subgoal_tac "finite (fname_length_set {fn. fn # f \<in> current_files \<tau>})")  
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    73
defer
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    74
apply (simp add:fname_length_set_def) 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    75
apply (rule finite_imageI)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    76
apply (drule_tac h = "\<lambda> f'. case f' of [] \<Rightarrow> '''' | fn # pf' \<Rightarrow> if (pf' = f) then fn else ''''" in finite_imageI)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    77
apply (rule_tac B = "(list_case [] (\<lambda>fn pf'. if pf' = f then fn else []) ` current_files \<tau>)" in finite_subset)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    78
unfolding image_def
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    79
apply (clarsimp  split:if_splits)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    80
apply (rule_tac x = "x # f" in bexI, simp+)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    81
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    82
apply (drule_tac s = "(fname_length_set {fn. fn # f \<in> current_files \<tau>})" in nn_notin_aux)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    83
apply (erule_tac x = "length (CHR ''a'' # fname_all_a (Max (fname_length_set {fn. fn # f \<in> current_files \<tau>})))" in ballE)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    84
apply (simp add:len_fname_all, simp)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    85
done
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    86
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    87
lemma ncf_parent: "valid \<tau> \<Longrightarrow> parent (new_childf f \<tau>) = Some f"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    88
by (simp add:new_childf_def)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    89
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    90
end
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    91
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    92
end