src/HOL/SET_Protocol/Event_SET.thy
author wenzelm
Mon, 10 Apr 2023 18:13:23 +0200
changeset 77804 849c996f052b
parent 69597 ff784d5a5bfb
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33028
9aa8bfb1649d modernized session SET_Protocol;
wenzelm
parents: 32960
diff changeset
     1
(*  Title:      HOL/SET_Protocol/Event_SET.thy
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30549
diff changeset
     2
    Author:     Giampaolo Bella
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30549
diff changeset
     3
    Author:     Fabio Massacci
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30549
diff changeset
     4
    Author:     Lawrence C Paulson
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
     5
*)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
     6
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 60754
diff changeset
     7
section\<open>Theory of Events for SET\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
     8
33028
9aa8bfb1649d modernized session SET_Protocol;
wenzelm
parents: 32960
diff changeset
     9
theory Event_SET
9aa8bfb1649d modernized session SET_Protocol;
wenzelm
parents: 32960
diff changeset
    10
imports Message_SET
9aa8bfb1649d modernized session SET_Protocol;
wenzelm
parents: 32960
diff changeset
    11
begin
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    12
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 60754
diff changeset
    13
text\<open>The Root Certification Authority\<close>
35101
6ce9177d6b38 modernized translations;
wenzelm
parents: 33028
diff changeset
    14
abbreviation "RCA == CA 0"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    15
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    16
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 60754
diff changeset
    17
text\<open>Message events\<close>
58310
91ea607a34d8 updated news
blanchet
parents: 58305
diff changeset
    18
datatype
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    19
  event = Says  agent agent msg
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30549
diff changeset
    20
        | Gets  agent       msg
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    21
        | Notes agent       msg
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    22
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    23
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 60754
diff changeset
    24
text\<open>compromised agents: keys known, Notes visible\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    25
consts bad :: "agent set"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    26
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 60754
diff changeset
    27
text\<open>Spy has access to his own key for spoof messages, but RCA is secure\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    28
specification (bad)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    29
  Spy_in_bad     [iff]: "Spy \<in> bad"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    30
  RCA_not_bad [iff]: "RCA \<notin> bad"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    31
    by (rule exI [of _ "{Spy}"], simp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    32
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    33
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 60754
diff changeset
    34
subsection\<open>Agents' Knowledge\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    35
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    36
consts  (*Initial states of agents -- parameter of the construction*)
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63167
diff changeset
    37
  initState :: "agent \<Rightarrow> msg set"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    38
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    39
(* Message reception does not extend spy's knowledge because of
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    40
   reception invariant enforced by Reception rule in protocol definition*)
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63167
diff changeset
    41
primrec knows :: "[agent, event list] \<Rightarrow> msg set"
39758
b8a53e3a0ee2 modernized primrecs
haftmann
parents: 35101
diff changeset
    42
where
b8a53e3a0ee2 modernized primrecs
haftmann
parents: 35101
diff changeset
    43
  knows_Nil:
b8a53e3a0ee2 modernized primrecs
haftmann
parents: 35101
diff changeset
    44
    "knows A [] = initState A"
b8a53e3a0ee2 modernized primrecs
haftmann
parents: 35101
diff changeset
    45
| knows_Cons:
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    46
    "knows A (ev # evs) =
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    47
       (if A = Spy then
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30549
diff changeset
    48
        (case ev of
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63167
diff changeset
    49
           Says A' B X \<Rightarrow> insert X (knows Spy evs)
ce654b0e6d69 more symbols;
wenzelm
parents: 63167
diff changeset
    50
         | Gets A' X \<Rightarrow> knows Spy evs
ce654b0e6d69 more symbols;
wenzelm
parents: 63167
diff changeset
    51
         | Notes A' X  \<Rightarrow>
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30549
diff changeset
    52
             if A' \<in> bad then insert X (knows Spy evs) else knows Spy evs)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30549
diff changeset
    53
        else
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30549
diff changeset
    54
        (case ev of
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63167
diff changeset
    55
           Says A' B X \<Rightarrow>
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30549
diff changeset
    56
             if A'=A then insert X (knows A evs) else knows A evs
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63167
diff changeset
    57
         | Gets A' X    \<Rightarrow>
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30549
diff changeset
    58
             if A'=A then insert X (knows A evs) else knows A evs
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63167
diff changeset
    59
         | Notes A' X    \<Rightarrow>
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30549
diff changeset
    60
             if A'=A then insert X (knows A evs) else knows A evs))"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    61
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    62
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 60754
diff changeset
    63
subsection\<open>Used Messages\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    64
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63167
diff changeset
    65
primrec used :: "event list \<Rightarrow> msg set"
39758
b8a53e3a0ee2 modernized primrecs
haftmann
parents: 35101
diff changeset
    66
where
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    67
  (*Set of items that might be visible to somebody:
39758
b8a53e3a0ee2 modernized primrecs
haftmann
parents: 35101
diff changeset
    68
    complement of the set of fresh items.
b8a53e3a0ee2 modernized primrecs
haftmann
parents: 35101
diff changeset
    69
    As above, message reception does extend used items *)
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    70
  used_Nil:  "used []         = (UN B. parts (initState B))"
39758
b8a53e3a0ee2 modernized primrecs
haftmann
parents: 35101
diff changeset
    71
| used_Cons: "used (ev # evs) =
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30549
diff changeset
    72
                 (case ev of
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63167
diff changeset
    73
                    Says A B X \<Rightarrow> parts {X} \<union> (used evs)
ce654b0e6d69 more symbols;
wenzelm
parents: 63167
diff changeset
    74
                  | Gets A X   \<Rightarrow> used evs
ce654b0e6d69 more symbols;
wenzelm
parents: 63167
diff changeset
    75
                  | Notes A X  \<Rightarrow> parts {X} \<union> (used evs))"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    76
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    77
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    78
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    79
(* Inserted by default but later removed.  This declaration lets the file
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    80
be re-loaded. Addsimps [knows_Cons, used_Nil, *)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    81
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    82
(** Simplifying   parts (insert X (knows Spy evs))
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63167
diff changeset
    83
      = parts {X} \<union> parts (knows Spy evs) -- since general case loops*)
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    84
45605
a89b4bc311a5 eliminated obsolete "standard";
wenzelm
parents: 39758
diff changeset
    85
lemmas parts_insert_knows_A = parts_insert [of _ "knows A evs"] for A evs
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    86
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    87
lemma knows_Spy_Says [simp]:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    88
     "knows Spy (Says A B X # evs) = insert X (knows Spy evs)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    89
by auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    90
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 60754
diff changeset
    91
text\<open>Letting the Spy see "bad" agents' notes avoids redundant case-splits
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
    92
      on whether \<^term>\<open>A=Spy\<close> and whether \<^term>\<open>A\<in>bad\<close>\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    93
lemma knows_Spy_Notes [simp]:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    94
     "knows Spy (Notes A X # evs) =
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63167
diff changeset
    95
          (if A\<in>bad then insert X (knows Spy evs) else knows Spy evs)"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    96
apply auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    97
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    98
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    99
lemma knows_Spy_Gets [simp]: "knows Spy (Gets A X # evs) = knows Spy evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   100
by auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   101
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63167
diff changeset
   102
lemma initState_subset_knows: "initState A \<subseteq> knows A evs"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   103
apply (induct_tac "evs")
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   104
apply (auto split: event.split) 
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   105
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   106
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   107
lemma knows_Spy_subset_knows_Spy_Says:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63167
diff changeset
   108
     "knows Spy evs \<subseteq> knows Spy (Says A B X # evs)"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   109
by auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   110
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   111
lemma knows_Spy_subset_knows_Spy_Notes:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63167
diff changeset
   112
     "knows Spy evs \<subseteq> knows Spy (Notes A X # evs)"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   113
by auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   114
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   115
lemma knows_Spy_subset_knows_Spy_Gets:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63167
diff changeset
   116
     "knows Spy evs \<subseteq> knows Spy (Gets A X # evs)"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   117
by auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   118
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   119
(*Spy sees what is sent on the traffic*)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   120
lemma Says_imp_knows_Spy [rule_format]:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63167
diff changeset
   121
     "Says A B X \<in> set evs \<longrightarrow> X \<in> knows Spy evs"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   122
apply (induct_tac "evs")
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   123
apply (auto split: event.split) 
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   124
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   125
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   126
(*Use with addSEs to derive contradictions from old Says events containing
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   127
  items known to be fresh*)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   128
lemmas knows_Spy_partsEs =
46471
2289a3869c88 prefer high-level elim_format;
wenzelm
parents: 45605
diff changeset
   129
     Says_imp_knows_Spy [THEN parts.Inj, elim_format] 
2289a3869c88 prefer high-level elim_format;
wenzelm
parents: 45605
diff changeset
   130
     parts.Body [elim_format]
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   131
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   132
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   133
subsection\<open>The Function \<^term>\<open>used\<close>\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   134
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63167
diff changeset
   135
lemma parts_knows_Spy_subset_used: "parts (knows Spy evs) \<subseteq> used evs"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   136
apply (induct_tac "evs")
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   137
apply (auto simp add: parts_insert_knows_A split: event.split) 
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   138
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   139
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   140
lemmas usedI = parts_knows_Spy_subset_used [THEN subsetD, intro]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   141
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63167
diff changeset
   142
lemma initState_subset_used: "parts (initState B) \<subseteq> used evs"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   143
apply (induct_tac "evs")
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   144
apply (auto split: event.split) 
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   145
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   146
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   147
lemmas initState_into_used = initState_subset_used [THEN subsetD]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   148
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63167
diff changeset
   149
lemma used_Says [simp]: "used (Says A B X # evs) = parts{X} \<union> used evs"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   150
by auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   151
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63167
diff changeset
   152
lemma used_Notes [simp]: "used (Notes A X # evs) = parts{X} \<union> used evs"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   153
by auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   154
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   155
lemma used_Gets [simp]: "used (Gets A X # evs) = used evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   156
by auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   157
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   158
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   159
lemma Notes_imp_parts_subset_used [rule_format]:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63167
diff changeset
   160
     "Notes A X \<in> set evs \<longrightarrow> parts {X} \<subseteq> used evs"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   161
apply (induct_tac "evs")
55417
01fbfb60c33e adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
blanchet
parents: 46471
diff changeset
   162
apply (rename_tac [2] a evs')
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   163
apply (induct_tac [2] "a", auto)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   164
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   165
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 60754
diff changeset
   166
text\<open>NOTE REMOVAL--laws above are cleaner, as they don't involve "case"\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   167
declare knows_Cons [simp del]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   168
        used_Nil [simp del] used_Cons [simp del]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   169
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   170
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   171
text\<open>For proving theorems of the form \<^term>\<open>X \<notin> analz (knows Spy evs) \<longrightarrow> P\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   172
  New events added by induction to "evs" are discarded.  Provided 
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   173
  this information isn't needed, the proof will be much shorter, since
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   174
  it will omit complicated reasoning about \<^term>\<open>analz\<close>.\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   175
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   176
lemmas analz_mono_contra =
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   177
       knows_Spy_subset_knows_Spy_Says [THEN analz_mono, THEN contra_subsetD]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   178
       knows_Spy_subset_knows_Spy_Notes [THEN analz_mono, THEN contra_subsetD]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   179
       knows_Spy_subset_knows_Spy_Gets [THEN analz_mono, THEN contra_subsetD]
27225
b316dde851f5 eliminated OldGoals.inst;
wenzelm
parents: 27159
diff changeset
   180
45605
a89b4bc311a5 eliminated obsolete "standard";
wenzelm
parents: 39758
diff changeset
   181
lemmas analz_impI = impI [where P = "Y \<notin> analz (knows Spy evs)"] for Y evs
27225
b316dde851f5 eliminated OldGoals.inst;
wenzelm
parents: 27159
diff changeset
   182
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   183
ML
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 60754
diff changeset
   184
\<open>
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58889
diff changeset
   185
fun analz_mono_contra_tac ctxt = 
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59498
diff changeset
   186
  resolve_tac ctxt @{thms analz_impI} THEN' 
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58963
diff changeset
   187
  REPEAT1 o (dresolve_tac ctxt @{thms analz_mono_contra})
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58889
diff changeset
   188
  THEN' mp_tac ctxt
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 60754
diff changeset
   189
\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   190
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 60754
diff changeset
   191
method_setup analz_mono_contra = \<open>
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 60754
diff changeset
   192
    Scan.succeed (fn ctxt => SIMPLE_METHOD (REPEAT_FIRST (analz_mono_contra_tac ctxt)))\<close>
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63167
diff changeset
   193
    "for proving theorems of the form X \<notin> analz (knows Spy evs) \<longrightarrow> P"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   194
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   195
end