no_shm_selinux/Enrich.thy
changeset 80 c9cfc416fa2d
parent 79 c09fcbcdb871
child 81 1ac0c3031ed2
equal deleted inserted replaced
79:c09fcbcdb871 80:c9cfc416fa2d
    28 | "enrich_proc (Open p f flags fd opt # s) tp dp = (
    28 | "enrich_proc (Open p f flags fd opt # s) tp dp = (
    29      if (tp = p)
    29      if (tp = p)
    30      then Open dp f (remove_create_flag flags) fd opt # Open p f flags fd opt # (enrich_proc s tp dp)
    30      then Open dp f (remove_create_flag flags) fd opt # Open p f flags fd opt # (enrich_proc s tp dp)
    31      else Open p f flags fd opt # (enrich_proc s tp dp))"
    31      else Open p f flags fd opt # (enrich_proc s tp dp))"
    32 | "enrich_proc (CloseFd p fd # s) tp dp = (
    32 | "enrich_proc (CloseFd p fd # s) tp dp = (
    33      if (tp = p)
    33      if (tp = p \<and> fd \<in> proc_file_fds s p)
    34      then CloseFd dp fd # CloseFd p fd # (enrich_proc s tp dp)
    34      then CloseFd dp fd # CloseFd p fd # (enrich_proc s tp dp)
    35      else CloseFd p fd # (enrich_proc s tp dp))"
    35      else CloseFd p fd # (enrich_proc s tp dp))"
    36 (*
    36 (*
    37 | "enrich_proc (Attach p h flag # s) tp dp = (
    37 | "enrich_proc (Attach p h flag # s) tp dp = (
    38      if (tp = p)
    38      if (tp = p)
   159   qed
   159   qed
   160   hence "sectxts_of_fds s' p fds = sectxts_of_fds s p fds"
   160   hence "sectxts_of_fds s' p fds = sectxts_of_fds s p fds"
   161     by (simp add:sectxts_of_fds_def)
   161     by (simp add:sectxts_of_fds_def)
   162   thus ?thesis using grant
   162   thus ?thesis using grant
   163     by (simp add:inherit_fds_check_def)
   163     by (simp add:inherit_fds_check_def)
       
   164 qed
       
   165 
       
   166 
       
   167 lemma enrich_inherit_fds_check_dup:
       
   168   assumes grant: "inherit_fds_check s (up, nr, nt) p fds"  and vd: "valid s"
       
   169   and cfd2sfd: "\<forall> fd. fd \<in> proc_file_fds s p \<longrightarrow> cfd2sfd s' p' fd = cfd2sfd s p fd"
       
   170   and fd_in: "fds' \<subseteq> fds \<inter> proc_file_fds s p" 
       
   171   shows "inherit_fds_check s' (up, nr, nt) p' fds'"
       
   172 proof- 
       
   173   have "sectxts_of_fds s' p' fds' \<subseteq> sectxts_of_fds s p fds"
       
   174   proof-
       
   175     have "\<And> fd sfd. \<lbrakk>fd \<in> fds'; sectxt_of_obj s' (O_fd p' fd) = Some sfd\<rbrakk> 
       
   176                      \<Longrightarrow> \<exists> fd \<in> fds. sectxt_of_obj s (O_fd p fd) = Some sfd"
       
   177     proof-
       
   178       fix fd sfd
       
   179       assume fd_in_fds': "fd \<in> fds'"
       
   180         and sec: "sectxt_of_obj s' (O_fd p' fd) = Some sfd"
       
   181       from fd_in_fds' fd_in
       
   182       have fd_in_fds: "fd \<in> fds" and fd_in_cfds: "fd \<in> proc_file_fds s p" 
       
   183         by auto
       
   184       from fd_in_cfds obtain f where ffd: "file_of_proc_fd s p fd = Some f"
       
   185         by (auto simp:proc_file_fds_def)
       
   186       moreover have "flags_of_proc_fd s p fd \<noteq> None"
       
   187         using ffd vd by (auto dest:current_filefd_has_flags)
       
   188       moreover have "cf2sfile s f \<noteq> None"
       
   189         apply (rule notI)
       
   190         apply (drule current_file_has_sfile')
       
   191         using ffd
       
   192         by (auto simp:vd is_file_in_current dest:file_of_pfd_is_file)
       
   193       moreover have "sectxt_of_obj s (O_fd p fd) \<noteq> None"
       
   194         using fd_in_cfds vd
       
   195         apply (rule_tac notI)
       
   196         by (auto dest!:current_has_sec' file_fds_subset_pfds[where p = p] intro:vd)
       
   197       ultimately 
       
   198       have "sectxt_of_obj s (O_fd p fd) = Some sfd"
       
   199         using fd_in_cfds cfd2sfd sec
       
   200         apply (erule_tac x = fd in allE)
       
   201         apply (auto simp:cfd2sfd_def split:option.splits)
       
   202         done
       
   203       thus "\<exists> fd \<in> fds. sectxt_of_obj s (O_fd p fd) = Some sfd"
       
   204         using fd_in_fds
       
   205         by (rule_tac x = fd in bexI, auto)
       
   206     qed
       
   207     thus ?thesis by (auto simp:sectxts_of_fds_def)
       
   208   qed
       
   209   thus ?thesis using grant
       
   210     by (auto simp:inherit_fds_check_def inherit_fds_check_ctxt_def)
   164 qed
   211 qed
   165 
   212 
   166 lemma not_all_procs_cons:
   213 lemma not_all_procs_cons:
   167   "p \<notin> all_procs (e # s) \<Longrightarrow> p \<notin> all_procs s"
   214   "p \<notin> all_procs (e # s) \<Longrightarrow> p \<notin> all_procs s"
   168 by (case_tac e, auto)
   215 by (case_tac e, auto)
  1103 apply (frule vt_grant_os, frule vd_cons)
  1150 apply (frule vt_grant_os, frule vd_cons)
  1104 apply (case_tac a, auto simp:is_created_proc_def proc_file_fds_def
  1151 apply (case_tac a, auto simp:is_created_proc_def proc_file_fds_def
  1105   dest:not_all_procs_prop3 split:if_splits option.splits)
  1152   dest:not_all_procs_prop3 split:if_splits option.splits)
  1106 done 
  1153 done 
  1107 
  1154 
       
  1155 lemma enrich_proc_dup_ffd':
       
  1156   "\<lbrakk>file_of_proc_fd (enrich_proc s p p') p' fd = Some f; is_created_proc s p; p' \<notin> all_procs s; valid s\<rbrakk>
       
  1157    \<Longrightarrow> file_of_proc_fd s p fd = Some f"
       
  1158 apply (induct s, simp add:is_created_proc_def)
       
  1159 apply (frule vt_grant_os, frule vd_cons)
       
  1160 apply (case_tac a, auto simp:is_created_proc_def proc_file_fds_def
       
  1161   dest:not_all_procs_prop3 split:if_splits option.splits)
       
  1162 done 
       
  1163 
  1108 lemma current_fflag_in_fds:
  1164 lemma current_fflag_in_fds:
  1109   "\<lbrakk>flags_of_proc_fd s p fd = Some flag; valid s\<rbrakk> \<Longrightarrow> fd \<in> current_proc_fds s p"
  1165   "\<lbrakk>flags_of_proc_fd s p fd = Some flag; valid s\<rbrakk> \<Longrightarrow> fd \<in> current_proc_fds s p"
  1110 apply (induct s arbitrary:p)
  1166 apply (induct s arbitrary:p)
  1111 apply (simp add:flags_of_proc_fd.simps file_of_proc_fd.simps init_oflags_prop2) 
  1167 apply (simp add:flags_of_proc_fd.simps file_of_proc_fd.simps init_oflags_prop2) 
  1112 apply (frule vd_cons, frule vt_grant_os, case_tac a)
  1168 apply (frule vd_cons, frule vt_grant_os, case_tac a)
  1113 apply (auto split:if_splits option.splits dest:proc_fd_in_fds)
  1169 apply (auto split:if_splits option.splits dest:proc_fd_in_fds)
  1114 apply (auto simp:proc_file_fds_def)
  1170 done
  1115 sorry
       
  1116 
       
  1117 
  1171 
  1118 lemma current_fflag_has_ffd:
  1172 lemma current_fflag_has_ffd:
  1119   "\<lbrakk>flags_of_proc_fd s p fd = Some flag; valid s\<rbrakk> \<Longrightarrow> \<exists> f. file_of_proc_fd s p fd = Some f"
  1173   "\<lbrakk>flags_of_proc_fd s p fd = Some flag; valid s\<rbrakk> \<Longrightarrow> \<exists> f. file_of_proc_fd s p fd = Some f"
  1120 apply (induct s arbitrary:p)
  1174 apply (induct s arbitrary:p)
  1121 apply (simp add: file_of_proc_fd.simps init_fileflag_valid) 
  1175 apply (simp add: file_of_proc_fd.simps init_fileflag_valid) 
  1122 apply (frule vd_cons, frule vt_grant_os, case_tac a)
  1176 apply (frule vd_cons, frule vt_grant_os, case_tac a)
  1123 apply (auto split:if_splits option.splits dest:proc_fd_in_fds)
  1177 apply (auto split:if_splits option.splits dest:proc_fd_in_fds)
       
  1178 done
       
  1179 
       
  1180 lemma enrich_proc_dup_fflags:
       
  1181   "\<lbrakk>flags_of_proc_fd s p fd = Some flag; is_created_proc s p; p' \<notin> all_procs s; valid s\<rbrakk>
       
  1182    \<Longrightarrow> flags_of_proc_fd (enrich_proc s p p') p' fd = Some (remove_create_flag flag) \<or>
       
  1183        flags_of_proc_fd (enrich_proc s p p') p' fd = Some flag"
       
  1184 apply (induct s arbitrary:p, simp add:is_created_proc_def)
       
  1185 apply (frule vt_grant_os, frule vd_cons)
       
  1186 apply (case_tac a, auto simp:is_created_proc_def proc_file_fds_def is_creat_flag_def
       
  1187   dest:not_all_procs_prop3 split:if_splits option.splits)
       
  1188 done
       
  1189 
       
  1190 lemma enrich_proc_dup_ffds:
       
  1191   "\<lbrakk>is_created_proc s p; p' \<notin> all_procs s; valid s\<rbrakk>
       
  1192    \<Longrightarrow> proc_file_fds (enrich_proc s p p') p' = proc_file_fds s p"
  1124 apply (auto simp:proc_file_fds_def)
  1193 apply (auto simp:proc_file_fds_def)
  1125 sorry
  1194 apply (rule_tac x = f in exI) 
  1126 
  1195 apply (erule enrich_proc_dup_ffd', simp+)
  1127 lemma enrich_proc_dup_fflags:
  1196 apply (rule_tac x = f in exI)
  1128   "\<lbrakk>flags_of_proc_fd s p fd = Some f; is_created_proc s p; p' \<notin> all_procs s; valid s\<rbrakk>
  1197 apply (erule enrich_proc_dup_ffd, simp+)
  1129    \<Longrightarrow> flags_of_proc_fd (enrich_proc s p p') p' fd = Some f"
  1198 done
  1130 apply (induct s, simp add:is_created_proc_def)
       
  1131 apply (frule vt_grant_os, frule vd_cons)
       
  1132 apply (case_tac a, auto simp:is_created_proc_def proc_file_fds_def
       
  1133   dest:not_all_procs_prop3 split:if_splits option.splits)
       
  1134 sorry
       
  1135 
  1199 
  1136 lemma enrich_proc_prop:
  1200 lemma enrich_proc_prop:
  1137   "\<lbrakk>valid s; is_created_proc s p; p' \<notin> all_procs s\<rbrakk>
  1201   "\<lbrakk>valid s; is_created_proc s p; p' \<notin> all_procs s\<rbrakk>
  1138    \<Longrightarrow> valid (enrich_proc s p p') \<and> 
  1202    \<Longrightarrow> valid (enrich_proc s p p') \<and> 
  1139        (\<forall> obj. alive s obj \<longrightarrow> alive (enrich_proc s p p') obj) \<and> 
  1203        (\<forall> obj. alive s obj \<longrightarrow> alive (enrich_proc s p p') obj) \<and> 
  1168      (\<forall>q. q \<in> current_msgqs s \<longrightarrow> cq2smsgq (enrich_proc s p p') q = cq2smsgq s q) \<and>
  1232      (\<forall>q. q \<in> current_msgqs s \<longrightarrow> cq2smsgq (enrich_proc s p p') q = cq2smsgq s q) \<and>
  1169      (\<forall>tp fd f. file_of_proc_fd s tp fd = Some f \<longrightarrow> file_of_proc_fd (enrich_proc s p p') tp fd = Some f) \<and>
  1233      (\<forall>tp fd f. file_of_proc_fd s tp fd = Some f \<longrightarrow> file_of_proc_fd (enrich_proc s p p') tp fd = Some f) \<and>
  1170      (\<forall>tp fd flags.
  1234      (\<forall>tp fd flags.
  1171          flags_of_proc_fd s tp fd = Some flags \<longrightarrow> flags_of_proc_fd (enrich_proc s p p') tp fd = Some flags) \<and>
  1235          flags_of_proc_fd s tp fd = Some flags \<longrightarrow> flags_of_proc_fd (enrich_proc s p p') tp fd = Some flags) \<and>
  1172      (\<forall>q. msgs_of_queue (enrich_proc s p p') q = msgs_of_queue s q) \<and>
  1236      (\<forall>q. msgs_of_queue (enrich_proc s p p') q = msgs_of_queue s q) \<and>
  1173      (\<forall>tp fd. fd \<in> proc_file_fds s tp \<longrightarrow> cfd2sfd (enrich_proc s p p') tp fd = cfd2sfd s tp fd)"
  1237      (\<forall>tp fd. fd \<in> proc_file_fds s tp \<longrightarrow> cfd2sfd (enrich_proc s p p') tp fd = cfd2sfd s tp fd) \<and>     
       
  1238      (cp2sproc (enrich_proc s p p') p' = cp2sproc s p) \<and>
       
  1239      (\<forall> fd. fd \<in> proc_file_fds s p \<longrightarrow> cfd2sfd (enrich_proc s p p') p' fd = cfd2sfd s p fd)"
  1174     using vd all_procs by auto
  1240     using vd all_procs by auto
       
  1241   have alive_pre: "is_created_proc s p \<Longrightarrow> (\<forall>obj. alive s obj \<longrightarrow> alive (enrich_proc s p p') obj)"
       
  1242     using pre by simp
       
  1243   hence curf_pre: "is_created_proc s p \<Longrightarrow> (\<forall>f. f \<in> current_files s \<longrightarrow> f \<in> current_files (enrich_proc s p p'))"
       
  1244     using vd apply auto
       
  1245     apply (drule is_file_or_dir, simp)
       
  1246     apply (erule disjE)
       
  1247     apply (erule_tac x = "O_file f" in allE, simp add:is_file_in_current)
       
  1248     apply (erule_tac x = "O_dir f" in allE, simp add:is_dir_in_current)
       
  1249     done
  1175   have "valid (enrich_proc (e # s) p p')"
  1250   have "valid (enrich_proc (e # s) p p')"
  1176   proof-
  1251   proof-
  1177     from pre have pre': "is_created_proc s p \<Longrightarrow> valid (enrich_proc s p p')" by simp
  1252     from pre have pre': "is_created_proc s p \<Longrightarrow> valid (enrich_proc s p p')" by simp
  1178     have "is_created_proc s p \<Longrightarrow> valid (e # enrich_proc s p p')" 
  1253     have "is_created_proc s p \<Longrightarrow> valid (e # enrich_proc s p p')" 
  1179       apply (frule pre')
  1254       apply (frule pre')
  1180       apply (erule_tac s = s in enrich_valid_intro_cons)
  1255       apply (erule_tac s = s in enrich_valid_intro_cons)
  1181       apply (simp_all add:os grant vd pre)
  1256       apply (simp_all add:os grant vd pre)
  1182       done  
  1257       done  
  1183     moreover have "\<And>f fds. \<lbrakk>valid (Execve p f fds # enrich_proc s p p'); is_created_proc s p\<rbrakk>
  1258     moreover have "\<And>f fds. \<lbrakk>valid (Execve p f fds # enrich_proc s p p'); is_created_proc s p; 
       
  1259       valid (Execve p f fds # s); p' \<notin> all_procs s\<rbrakk>
  1184       \<Longrightarrow> valid (Execve p' f (fds \<inter> proc_file_fds s p) # Execve p f fds # enrich_proc s p p')"
  1260       \<Longrightarrow> valid (Execve p' f (fds \<inter> proc_file_fds s p) # Execve p f fds # enrich_proc s p p')"
  1185       apply (frule vt_grant_os, frule vt_grant)
  1261     proof-
  1186       apply (rule valid.intros(2))
  1262       fix f fds
  1187       apply (simp_all split:option.splits add:sectxt_of_obj_simps is_file_simps)
  1263       assume a1: "valid (Execve p f fds # enrich_proc s p p')" and a2: "is_created_proc s p"
  1188       apply (auto simp add:proc_file_fds_def)[1]
  1264         and a3: "valid (Execve p f fds # s)" and a0: "p' \<notin> all_procs s"
  1189       
  1265       have cp2sp: "\<forall> tp. tp \<in> current_procs s \<longrightarrow> cp2sproc (enrich_proc s p p') tp = cp2sproc s tp"
  1190       sorry
  1266         and cf2sf: "\<forall> tf. tf \<in> current_files s \<longrightarrow> cf2sfile (enrich_proc s p p') tf = cf2sfile s tf"
       
  1267         and cfd2sfd: "\<forall> tp tfd. tfd \<in> proc_file_fds s tp \<longrightarrow> cfd2sfd (enrich_proc s p p') tp tfd = cfd2sfd s tp tfd"
       
  1268         and ffd_remain: "\<forall>tp fd f. file_of_proc_fd s tp fd = Some f \<longrightarrow> 
       
  1269                                    file_of_proc_fd (enrich_proc s p p') tp fd = Some f"
       
  1270         and dup_sp: "cp2sproc (enrich_proc s p p') p' = cp2sproc s p"
       
  1271         and dup_sfd: "\<forall> fd. fd \<in> proc_file_fds s p \<longrightarrow> cfd2sfd (enrich_proc s p p') p' fd = cfd2sfd s p fd"
       
  1272         using pre a2 by auto
       
  1273       show "valid (Execve p' f (fds \<inter> proc_file_fds s p) # Execve p f fds # enrich_proc s p p')"
       
  1274       proof-
       
  1275         from a0 a3 have a0': "p' \<noteq> p" by (auto dest!:vt_grant_os not_all_procs_prop3)
       
  1276         from a3 have grant: "grant s (Execve p f fds)" and os: "os_grant s (Execve p f fds)"
       
  1277           by (auto dest:vt_grant_os vt_grant simp del:os_grant.simps)
       
  1278         have f_in: "is_file (enrich_proc s p p') f" 
       
  1279         proof-
       
  1280           from pre a2
       
  1281           have a4: "\<forall> obj. alive s obj \<longrightarrow> alive (enrich_proc s p p') obj"
       
  1282             by (auto)
       
  1283           show ?thesis using a3 a4
       
  1284             apply (erule_tac x = "O_file f" in allE)
       
  1285             by (auto dest:vt_grant_os)
       
  1286         qed
       
  1287         moreover have a5: "proc_file_fds s p \<subseteq> proc_file_fds (Execve p f fds # enrich_proc s p p') p'" 
       
  1288           using a3 a0'
       
  1289           apply (frule_tac vt_grant_os)
       
  1290           apply (auto simp:proc_file_fds_def)
       
  1291           apply (rule_tac x = fa in exI)
       
  1292           apply (erule enrich_proc_dup_ffd)
       
  1293           apply (simp_all add:vd all_procs a2)
       
  1294           done
       
  1295         ultimately have "os_grant (Execve p f fds # enrich_proc s p p') (Execve p' f (fds \<inter> proc_file_fds s p))"
       
  1296           apply (auto simp:is_file_simps enrich_proc_dup_in a2 vd all_procs a1 enrich_proc_dup_ffds)
       
  1297           done
       
  1298         moreover have "grant (Execve p f fds # enrich_proc s p p') (Execve p' f (fds \<inter> proc_file_fds s p))"
       
  1299         proof-
       
  1300           from grant obtain up rp tp uf rf tf 
       
  1301             where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)"
       
  1302             and p2: "sectxt_of_obj s (O_file f) = Some (uf, rf, tf)" 
       
  1303             by (simp split:option.splits, blast)
       
  1304           with grant obtain pu nr nt where p3: "npctxt_execve (up, rp, tp) (uf, rf, tf) = Some (pu, nr, nt)"
       
  1305             by (simp split:option.splits del:npctxt_execve.simps, blast)
       
  1306           have p1': "sectxt_of_obj (Execve p f fds # enrich_proc s p p') (O_proc p') = Some (up, rp, tp)"
       
  1307             using p1 dup_sp a1 a0'
       
  1308             apply (simp add:sectxt_of_obj_simps)
       
  1309             by (simp add:cp2sproc_def split:option.splits)
       
  1310           from os have f_in': "is_file s f"  by simp 
       
  1311           from vd os have "\<exists> sf. cf2sfile s f = Some sf"
       
  1312             by (auto dest!:is_file_in_current current_file_has_sfile)
       
  1313           hence p2': "sectxt_of_obj (Execve p f fds # enrich_proc s p p') (O_file f) = Some (uf, rf, tf)" 
       
  1314             using f_in p2 cf2sf os a1
       
  1315             apply (erule_tac x = f in allE)
       
  1316             apply (auto dest:is_file_in_current simp:cf2sfile_def sectxt_of_obj_simps split:option.splits)
       
  1317             apply (case_tac f, simp)
       
  1318             apply (drule_tac s = s in root_is_dir', simp add:vd, simp+)
       
  1319             done 
       
  1320           from dup_sfd a5 have "\<forall>fd. fd \<in> proc_file_fds s p \<longrightarrow> 
       
  1321             cfd2sfd (Execve p f fds # enrich_proc s p p') p' fd = cfd2sfd s p fd"
       
  1322             apply (rule_tac allI)
       
  1323             apply (erule_tac x = fd in allE, clarsimp)
       
  1324             apply (drule set_mp, simp)
       
  1325             apply (auto simp:cfd2sfd_execve proc_file_fds_def a1)
       
  1326             done
       
  1327           hence "inherit_fds_check (Execve p f fds # enrich_proc s p p') (up, nr, nt) p' (fds \<inter> proc_file_fds s p)"
       
  1328             using grant os p1 p2 p3 vd
       
  1329             apply (clarsimp)
       
  1330             apply (rule_tac s = s and p = p and fds = fds in enrich_inherit_fds_check_dup)
       
  1331             apply simp_all
       
  1332             done
       
  1333           moreover have "search_check (Execve p f fds # enrich_proc s p p') (up, rp, tp) f"
       
  1334             using p1 p2 p2' vd cf2sf f_in f_in' grant p3 f_in a1
       
  1335             apply (rule_tac s = s in enrich_search_check)
       
  1336             apply (simp_all add:is_file_simps)
       
  1337             apply (rule allI, rule impI, erule_tac x = fa in allE, simp)
       
  1338             apply (drule_tac ff = fa in cf2sfile_other')
       
  1339             by (auto simp:a2 curf_pre)
       
  1340           ultimately show ?thesis 
       
  1341             using p1' p2' p3
       
  1342             apply (simp split:option.splits)
       
  1343             using grant p1 p2
       
  1344             apply simp
       
  1345             done
       
  1346         qed
       
  1347         ultimately show ?thesis using a1
       
  1348           by (erule_tac valid.intros(2), auto)
       
  1349       qed
       
  1350     qed
  1191     moreover have "\<And>tp fds. \<lbrakk>valid (Clone tp p fds # s); p' \<noteq> p; p' \<notin> all_procs s\<rbrakk> \<Longrightarrow> 
  1351     moreover have "\<And>tp fds. \<lbrakk>valid (Clone tp p fds # s); p' \<noteq> p; p' \<notin> all_procs s\<rbrakk> \<Longrightarrow> 
  1192       valid (Clone tp p' (fds \<inter> proc_file_fds s tp) # Clone tp p fds # s)"
  1352       valid (Clone tp p' (fds \<inter> proc_file_fds s tp) # Clone tp p fds # s)"
  1193       apply (frule vt_grant_os, frule vt_grant, drule not_all_procs_prop3)
  1353       apply (frule vt_grant_os, frule vt_grant, drule not_all_procs_prop3)
  1194       apply (rule valid.intros(2))
  1354       apply (rule valid.intros(2))
  1195       apply (simp_all split:option.splits add:sectxt_of_obj_simps)
  1355       apply (simp_all split:option.splits add:sectxt_of_obj_simps)
  1198       done
  1358       done
  1199     moreover have "\<And>f flags fd inum.
  1359     moreover have "\<And>f flags fd inum.
  1200        \<lbrakk>valid (Open p f flags fd inum # enrich_proc s p p'); is_created_proc s p\<rbrakk>
  1360        \<lbrakk>valid (Open p f flags fd inum # enrich_proc s p p'); is_created_proc s p\<rbrakk>
  1201        \<Longrightarrow> valid (Open p' f (remove_create_flag flags) fd inum # Open p f flags fd inum # enrich_proc s p p')"
  1361        \<Longrightarrow> valid (Open p' f (remove_create_flag flags) fd inum # Open p f flags fd inum # enrich_proc s p p')"
  1202       sorry
  1362       sorry
  1203     moreover have "\<And>fd. \<lbrakk>valid (CloseFd p fd # enrich_proc s p p'); is_created_proc s p\<rbrakk>
  1363     moreover have "\<And>fd. \<lbrakk>valid (CloseFd p fd # enrich_proc s p p'); is_created_proc s p; 
       
  1364       valid (CloseFd p fd # s); p' \<notin> all_procs s\<rbrakk>
  1204            \<Longrightarrow> valid (CloseFd p' fd # CloseFd p fd # enrich_proc s p p')"
  1365            \<Longrightarrow> valid (CloseFd p' fd # CloseFd p fd # enrich_proc s p p')"
       
  1366     proof-
       
  1367       fix fd
       
  1368       assume a1: "valid (CloseFd p fd # enrich_proc s p p')" and a2: "is_created_proc s p" 
       
  1369         and a3: "p' \<notin> all_procs s" and a4: "valid (CloseFd p fd # s)"
       
  1370       from a4 a3 have a0: "p' \<noteq> p" by (auto dest!:vt_grant_os not_all_procs_prop3)
       
  1371       have "p' \<in> current_procs (enrich_proc s p p')" 
       
  1372         using a2 a3 vd
       
  1373         by (auto intro:enrich_proc_dup_in)
       
  1374       moreover have "fd \<in> current_proc_fds (enrich_proc s p p') p'"
       
  1375       ultimately show "valid (CloseFd p' fd # CloseFd p fd # enrich_proc s p p')"
       
  1376         apply (rule valid.intros(2))
       
  1377         apply (simp_all add:a1 a0 a2 pre')
       
  1378       
       
  1379 
       
  1380 
  1205       sorry 
  1381       sorry 
  1206         thus ?thesis using created_cons vd_cons all_procs_cons
  1382 (*
       
  1383        thus ?thesis using created_cons vd_cons all_procs_cons
  1207       apply (case_tac e)
  1384       apply (case_tac e)
  1208       apply (auto simp:is_created_proc_simps split:if_splits)
  1385       apply (auto simp:is_created_proc_simps split:if_splits)
  1209         
  1386         *)
  1210     ultimately show ?thesis using created_cons vd_cons all_procs_cons
  1387     ultimately show ?thesis using created_cons vd_cons all_procs_cons
  1211       apply (case_tac e)
  1388       apply (case_tac e)
  1212       apply (auto simp:is_created_proc_simps split:if_splits)
  1389       apply (auto simp:is_created_proc_simps split:if_splits)
  1213       done
  1390       done
  1214   qed
  1391   qed