simple_selinux/Finite_current.thy
author chunhan
Wed, 08 Jan 2014 18:40:38 +0800
changeset 91 1a1df29d3507
parent 74 271e9818b6f6
permissions -rw-r--r--
enrich msgq done; but find bugs of s2ss, it should only considerate 'appropriate' objects, not including msg/fd ...
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
74
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
     1
theory Finite_current
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
     2
imports Main Valid_prop Flask Flask_type Proc_fd_of_file_prop
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
     3
begin
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
     4
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
     5
context flask begin
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
     6
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
     7
lemma finite_cf: "valid \<tau> \<Longrightarrow> finite (current_files \<tau>)"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
     8
apply (induct \<tau>)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
     9
apply (simp add:current_files_def inum_of_file.simps)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    10
apply (rule_tac B = "init_files" in finite_subset)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    11
apply (clarsimp dest!:inof_has_file_tag, simp add:init_finite_sets)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    12
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    13
apply (frule vt_grant_os, frule vd_cons, simp, case_tac a)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    14
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    15
apply (auto simp:current_files_def os_grant.simps inum_of_file.simps split:if_splits option.splits)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    16
apply (rule_tac B = "insert list {f. \<exists>i. inum_of_file \<tau> f = Some i}" in finite_subset, clarsimp, simp)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    17
apply (rule_tac B = "{f. \<exists>i. inum_of_file \<tau> f = Some i}" in finite_subset, clarsimp, simp)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    18
apply (rule_tac B = "{f. \<exists>i. inum_of_file \<tau> f = Some i}" in finite_subset, clarsimp, simp)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    19
apply (rule_tac B = "{f. \<exists>i. inum_of_file \<tau> f = Some i}" in finite_subset, clarsimp, simp)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    20
apply (rule_tac B = "insert list {f. \<exists>i. inum_of_file \<tau> f = Some i}" in finite_subset, clarsimp, simp)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    21
done
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    22
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    23
lemma finite_cp: "finite (current_procs \<tau>)"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    24
apply (induct \<tau>)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    25
apply (simp add:current_procs.simps init_finite_sets)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    26
apply (case_tac a, auto simp:current_procs.simps)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    27
done
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    28
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    29
lemma finite_cfd: "valid \<tau> \<Longrightarrow> finite (current_proc_fds \<tau> p)"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    30
apply (induct \<tau> arbitrary:p)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    31
apply (simp add:current_proc_fds.simps init_finite_sets)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    32
apply (frule vd_cons, frule vt_grant_os, case_tac a, auto simp:current_proc_fds.simps)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    33
apply (erule finite_subset)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    34
apply (frule_tac s = \<tau> and p = nat in file_fds_subset_pfds)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    35
apply (erule finite_subset, simp)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    36
apply (erule finite_subset)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    37
apply (frule_tac s = \<tau> and p = nat1 in file_fds_subset_pfds)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    38
apply (erule finite_subset, simp)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    39
done
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    40
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    41
lemma finite_pair: "\<lbrakk>finite A; finite B\<rbrakk> \<Longrightarrow> finite {(x, y). x \<in> A \<and> y \<in> B}"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    42
by auto
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    43
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    44
lemma finite_UN_I': "\<lbrakk>finite X; \<forall> x. x \<in> X \<longrightarrow> finite (f x)\<rbrakk> \<Longrightarrow> finite {(x, y). x \<in> X \<and> y \<in> f x}"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    45
apply (frule_tac B = f in finite_UN_I, simp)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    46
apply (drule_tac finite_pair, simp)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    47
apply (rule_tac B = "{(x, y). x \<in> X \<and> y \<in> (\<Union>a\<in>X. f a)}" in finite_subset, auto)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    48
done
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    49
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    50
lemma finite_init_netobjs: "finite init_sockets"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    51
apply (subgoal_tac "finite {(p, fd). p \<in> init_procs \<and> fd \<in> init_fds_of_proc p}")
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    52
apply (rule_tac B = "{(p, fd). p \<in> init_procs \<and> fd \<in> init_fds_of_proc p}" in finite_subset)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    53
apply (clarsimp dest!:init_socket_has_inode, simp)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    54
using init_finite_sets finite_UN_I'
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    55
by (metis Collect_mem_eq SetCompr_Sigma_eq internal_split_def) 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    56
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    57
lemma finite_cn_aux: "valid \<tau> \<Longrightarrow> finite {s. \<exists>i. inum_of_socket \<tau> s = Some i}"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    58
apply (induct \<tau>)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    59
apply (rule_tac B = "init_sockets" in finite_subset)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    60
apply (clarsimp simp:inum_of_socket.simps dest!:inos_has_sock_tag, simp add:finite_init_netobjs)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    61
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    62
apply (frule vd_cons, frule vt_grant_os, simp, case_tac a)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    63
apply (auto split:option.splits if_splits) 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    64
apply (rule_tac B = "{s. \<exists>i. inum_of_socket \<tau> s = Some i}" in finite_subset, clarsimp split:if_splits, simp)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    65
apply (rule_tac B = "{s. \<exists>i. inum_of_socket \<tau> s = Some i} \<union> {(p, fd). \<exists> i. inum_of_socket \<tau> (nat1, fd) = Some i \<and> p = nat2 \<and> fd \<in> set}" in finite_subset, clarsimp split:if_splits)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    66
apply (simp only:finite_Un, rule conjI, simp)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    67
apply (rule_tac B = "{(p, fd). \<exists> i. inum_of_socket \<tau> (nat1, fd) = Some i \<and> p = nat2}" in finite_subset, clarsimp)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    68
apply (drule_tac h = "\<lambda> (p, fd). if (p = nat1) then (nat2, fd) else (p, fd)" in finite_imageI)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    69
apply (rule_tac B = "((\<lambda>(p, fd). if p = nat1 then (nat2, fd) else (p, fd)) ` {a. \<exists>i. inum_of_socket \<tau> a = Some i})" in finite_subset) 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    70
apply (rule subsetI,erule CollectE, case_tac x, simp, (erule exE|erule conjE)+)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    71
unfolding image_def
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    72
apply (rule CollectI, rule_tac x = "(nat1, b)" in bexI, simp+)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    73
apply (rule_tac B = "{s. \<exists>i. inum_of_socket \<tau> s = Some i}" in finite_subset, clarsimp split:if_splits, simp)+
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    74
apply (rule_tac B = "insert (nat1, nat2) {s. \<exists>i. inum_of_socket \<tau> s = Some i}" in finite_subset, clarsimp, simp)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    75
apply (rule_tac B = "insert (nat1, nat4) {s. \<exists>i. inum_of_socket \<tau> s = Some i}" in finite_subset, clarsimp, simp)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    76
done
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    77
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    78
lemma finite_cn: "valid \<tau> \<Longrightarrow> finite (current_sockets \<tau>)"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    79
apply (simp add:current_sockets_def inum_of_socket.simps)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    80
using finite_cn_aux[where \<tau> = \<tau>] by auto
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    81
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    82
lemma finite_cm: "finite (current_msgqs \<tau>)"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    83
apply (induct \<tau>) defer
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    84
apply (case_tac a, auto simp: init_finite_sets)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    85
done
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    86
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    87
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    88
lemma finite_option: "finite {x. \<exists> y. f x = Some y} \<Longrightarrow> finite {y. \<exists> x. f x = Some y}"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    89
apply (drule_tac h = f in finite_imageI)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    90
apply (clarsimp simp only:image_def)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    91
apply (rule_tac f = Some in finite_imageD)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    92
apply (rule_tac B = "{y. \<exists>x. (\<exists>y. f x = Some y) \<and> y = f x}" in finite_subset)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    93
unfolding image_def
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    94
apply auto
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    95
done
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    96
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    97
lemma finite_ci: "valid \<tau> \<Longrightarrow> finite (current_inode_nums \<tau>)"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    98
apply (simp add:current_inode_nums_def current_file_inums_def current_sock_inums_def)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    99
apply (rule conjI, drule finite_cf, simp add:current_files_def, erule finite_option) 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   100
using finite_cn[where \<tau> = \<tau>] 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   101
apply (simp add:current_sockets_def, drule_tac finite_option, simp)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   102
done
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   103
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   104
end
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   105
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   106
end
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   107
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   108