deleted_prop.thy
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Sat, 14 Dec 2013 13:07:41 +1100
changeset 20 928c015eb03e
parent 1 dcde836219bc
permissions -rw-r--r--
updated slides
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1
dcde836219bc add thy files
chunhan
parents:
diff changeset
     1
theory deleted_prop
dcde836219bc add thy files
chunhan
parents:
diff changeset
     2
imports Main rc_theory os_rc
dcde836219bc add thy files
chunhan
parents:
diff changeset
     3
begin
dcde836219bc add thy files
chunhan
parents:
diff changeset
     4
dcde836219bc add thy files
chunhan
parents:
diff changeset
     5
context tainting begin
dcde836219bc add thy files
chunhan
parents:
diff changeset
     6
dcde836219bc add thy files
chunhan
parents:
diff changeset
     7
lemma deleted_cons_I: "deleted obj s \<Longrightarrow> deleted obj (e # s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
     8
by (case_tac e, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
     9
dcde836219bc add thy files
chunhan
parents:
diff changeset
    10
lemma not_deleted_cons_D: "\<not> deleted obj (e # s) \<Longrightarrow> \<not> deleted obj s" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    11
by (auto dest:deleted_cons_I)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    12
dcde836219bc add thy files
chunhan
parents:
diff changeset
    13
lemma not_deleted_imp_exists:
dcde836219bc add thy files
chunhan
parents:
diff changeset
    14
  "\<lbrakk>\<not> deleted obj s; exists [] obj\<rbrakk> \<Longrightarrow> exists s obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    15
apply (induct s, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    16
apply (case_tac a, case_tac [!] obj, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    17
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
    18
dcde836219bc add thy files
chunhan
parents:
diff changeset
    19
lemma cons_app_simp_aux:
dcde836219bc add thy files
chunhan
parents:
diff changeset
    20
  "(a # b) @ c = a # (b @ c)" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
    21
dcde836219bc add thy files
chunhan
parents:
diff changeset
    22
lemma not_deleted_imp_exists':
dcde836219bc add thy files
chunhan
parents:
diff changeset
    23
  "\<lbrakk>\<not> deleted obj (s'@s); exists s obj\<rbrakk> \<Longrightarrow> exists (s'@s) obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    24
apply (induct s', simp, simp only:cons_app_simp_aux)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    25
apply (frule not_deleted_cons_D)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    26
apply (case_tac a, case_tac [!] obj, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    27
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
    28
dcde836219bc add thy files
chunhan
parents:
diff changeset
    29
lemma nodel_imp_un_deleted:
dcde836219bc add thy files
chunhan
parents:
diff changeset
    30
  "no_del_event s \<Longrightarrow> \<not> deleted obj s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    31
by (induct s, simp, case_tac a,auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    32
dcde836219bc add thy files
chunhan
parents:
diff changeset
    33
lemma nodel_exists_remains:
dcde836219bc add thy files
chunhan
parents:
diff changeset
    34
  "\<lbrakk>no_del_event (s'@s); exists s obj\<rbrakk> \<Longrightarrow> exists (s'@s) obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    35
apply (drule_tac obj = obj in nodel_imp_un_deleted)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    36
by (simp add:not_deleted_imp_exists')
dcde836219bc add thy files
chunhan
parents:
diff changeset
    37
dcde836219bc add thy files
chunhan
parents:
diff changeset
    38
lemma nodel_imp_exists:
dcde836219bc add thy files
chunhan
parents:
diff changeset
    39
  "\<lbrakk>no_del_event s; exists [] obj\<rbrakk> \<Longrightarrow> exists s obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    40
apply (drule_tac obj = obj in nodel_imp_un_deleted)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    41
by (simp add:not_deleted_imp_exists)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    42
dcde836219bc add thy files
chunhan
parents:
diff changeset
    43
lemma no_del_event_cons_D:
dcde836219bc add thy files
chunhan
parents:
diff changeset
    44
  "no_del_event (e # s) \<Longrightarrow> no_del_event s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    45
by (case_tac e, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    46
dcde836219bc add thy files
chunhan
parents:
diff changeset
    47
end
dcde836219bc add thy files
chunhan
parents:
diff changeset
    48
dcde836219bc add thy files
chunhan
parents:
diff changeset
    49
end