src/HOL/Statespace/DistinctTreeProver.thy
author wenzelm
Wed, 14 May 2014 12:00:18 +0200
changeset 56958 b2c2f74d1c93
parent 55378 e61e023c9faf
child 58249 180f1b3508ed
permissions -rw-r--r--
updated to polyml-5.5.2;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
29269
5c25a2012975 moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
wenzelm
parents: 28819
diff changeset
     1
(*  Title:      HOL/Statespace/DistinctTreeProver.thy
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
     2
    Author:     Norbert Schirmer, TU Muenchen
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
     3
*)
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
     4
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
     5
header {* Distinctness of Names in a Binary Tree \label{sec:DistinctTreeProver}*}
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
     6
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
     7
theory DistinctTreeProver 
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
     8
imports Main
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
     9
begin
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    10
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    11
text {* A state space manages a set of (abstract) names and assumes
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    12
that the names are distinct. The names are stored as parameters of a
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    13
locale and distinctness as an assumption. The most common request is
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    14
to proof distinctness of two given names. We maintain the names in a
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    15
balanced binary tree and formulate a predicate that all nodes in the
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    16
tree have distinct names. This setup leads to logarithmic certificates.
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    17
*}
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    18
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    19
subsection {* The Binary Tree *}
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    20
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    21
datatype 'a tree = Node "'a tree" 'a bool "'a tree" | Tip
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    22
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    23
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    24
text {* The boolean flag in the node marks the content of the node as
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    25
deleted, without having to build a new tree. We prefer the boolean
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    26
flag to an option type, so that the ML-layer can still use the node
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    27
content to facilitate binary search in the tree. The ML code keeps the
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    28
nodes sorted using the term order. We do not have to push ordering to
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    29
the HOL level. *}
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    30
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    31
subsection {* Distinctness of Nodes *}
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    32
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    33
38838
62f6ba39b3d4 modernized specifications;
wenzelm
parents: 35408
diff changeset
    34
primrec set_of :: "'a tree \<Rightarrow> 'a set"
62f6ba39b3d4 modernized specifications;
wenzelm
parents: 35408
diff changeset
    35
where
62f6ba39b3d4 modernized specifications;
wenzelm
parents: 35408
diff changeset
    36
  "set_of Tip = {}"
62f6ba39b3d4 modernized specifications;
wenzelm
parents: 35408
diff changeset
    37
| "set_of (Node l x d r) = (if d then {} else {x}) \<union> set_of l \<union> set_of r"
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    38
38838
62f6ba39b3d4 modernized specifications;
wenzelm
parents: 35408
diff changeset
    39
primrec all_distinct :: "'a tree \<Rightarrow> bool"
62f6ba39b3d4 modernized specifications;
wenzelm
parents: 35408
diff changeset
    40
where
62f6ba39b3d4 modernized specifications;
wenzelm
parents: 35408
diff changeset
    41
  "all_distinct Tip = True"
62f6ba39b3d4 modernized specifications;
wenzelm
parents: 35408
diff changeset
    42
| "all_distinct (Node l x d r) =
62f6ba39b3d4 modernized specifications;
wenzelm
parents: 35408
diff changeset
    43
    ((d \<or> (x \<notin> set_of l \<and> x \<notin> set_of r)) \<and> 
62f6ba39b3d4 modernized specifications;
wenzelm
parents: 35408
diff changeset
    44
      set_of l \<inter> set_of r = {} \<and>
62f6ba39b3d4 modernized specifications;
wenzelm
parents: 35408
diff changeset
    45
      all_distinct l \<and> all_distinct r)"
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    46
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    47
text {* Given a binary tree @{term "t"} for which 
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    48
@{const all_distinct} holds, given two different nodes contained in the tree,
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    49
we want to write a ML function that generates a logarithmic
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    50
certificate that the content of the nodes is distinct. We use the
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    51
following lemmas to achieve this.  *} 
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    52
45355
c0704e988526 misc tuning and modernization;
wenzelm
parents: 44890
diff changeset
    53
lemma all_distinct_left: "all_distinct (Node l x b r) \<Longrightarrow> all_distinct l"
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    54
  by simp
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    55
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    56
lemma all_distinct_right: "all_distinct (Node l x b r) \<Longrightarrow> all_distinct r"
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    57
  by simp
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    58
45358
4849133d7a78 tuned document;
wenzelm
parents: 45355
diff changeset
    59
lemma distinct_left: "all_distinct (Node l x False r) \<Longrightarrow> y \<in> set_of l \<Longrightarrow> x \<noteq> y"
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    60
  by auto
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    61
45358
4849133d7a78 tuned document;
wenzelm
parents: 45355
diff changeset
    62
lemma distinct_right: "all_distinct (Node l x False r) \<Longrightarrow> y \<in> set_of r \<Longrightarrow> x \<noteq> y"
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    63
  by auto
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    64
45358
4849133d7a78 tuned document;
wenzelm
parents: 45355
diff changeset
    65
lemma distinct_left_right:
4849133d7a78 tuned document;
wenzelm
parents: 45355
diff changeset
    66
    "all_distinct (Node l z b r) \<Longrightarrow> x \<in> set_of l \<Longrightarrow> y \<in> set_of r \<Longrightarrow> x \<noteq> y"
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    67
  by auto
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    68
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    69
lemma in_set_root: "x \<in> set_of (Node l x False r)"
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    70
  by simp
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    71
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    72
lemma in_set_left: "y \<in> set_of l \<Longrightarrow>  y \<in> set_of (Node l x False r)"
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    73
  by simp
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    74
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    75
lemma in_set_right: "y \<in> set_of r \<Longrightarrow>  y \<in> set_of (Node l x False r)"
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    76
  by simp
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    77
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    78
lemma swap_neq: "x \<noteq> y \<Longrightarrow> y \<noteq> x"
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    79
  by blast
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    80
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    81
lemma neq_to_eq_False: "x\<noteq>y \<Longrightarrow> (x=y)\<equiv>False"
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    82
  by simp
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    83
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    84
subsection {* Containment of Trees *}
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    85
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    86
text {* When deriving a state space from other ones, we create a new
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    87
name tree which contains all the names of the parent state spaces and
45358
4849133d7a78 tuned document;
wenzelm
parents: 45355
diff changeset
    88
assume the predicate @{const all_distinct}. We then prove that the new
4849133d7a78 tuned document;
wenzelm
parents: 45355
diff changeset
    89
locale interprets all parent locales. Hence we have to show that the
4849133d7a78 tuned document;
wenzelm
parents: 45355
diff changeset
    90
new distinctness assumption on all names implies the distinctness
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    91
assumptions of the parent locales. This proof is implemented in ML. We
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    92
do this efficiently by defining a kind of containment check of trees
45358
4849133d7a78 tuned document;
wenzelm
parents: 45355
diff changeset
    93
by ``subtraction''.  We subtract the parent tree from the new tree. If
4849133d7a78 tuned document;
wenzelm
parents: 45355
diff changeset
    94
this succeeds we know that @{const all_distinct} of the new tree
4849133d7a78 tuned document;
wenzelm
parents: 45355
diff changeset
    95
implies @{const all_distinct} of the parent tree.  The resulting
4849133d7a78 tuned document;
wenzelm
parents: 45355
diff changeset
    96
certificate is of the order @{term "n * log(m)"} where @{term "n"} is
4849133d7a78 tuned document;
wenzelm
parents: 45355
diff changeset
    97
the size of the (smaller) parent tree and @{term "m"} the size of the
4849133d7a78 tuned document;
wenzelm
parents: 45355
diff changeset
    98
(bigger) new tree.  *}
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
    99
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   100
38838
62f6ba39b3d4 modernized specifications;
wenzelm
parents: 35408
diff changeset
   101
primrec delete :: "'a \<Rightarrow> 'a tree \<Rightarrow> 'a tree option"
62f6ba39b3d4 modernized specifications;
wenzelm
parents: 35408
diff changeset
   102
where
62f6ba39b3d4 modernized specifications;
wenzelm
parents: 35408
diff changeset
   103
  "delete x Tip = None"
62f6ba39b3d4 modernized specifications;
wenzelm
parents: 35408
diff changeset
   104
| "delete x (Node l y d r) = (case delete x l of
62f6ba39b3d4 modernized specifications;
wenzelm
parents: 35408
diff changeset
   105
                                Some l' \<Rightarrow>
62f6ba39b3d4 modernized specifications;
wenzelm
parents: 35408
diff changeset
   106
                                 (case delete x r of 
62f6ba39b3d4 modernized specifications;
wenzelm
parents: 35408
diff changeset
   107
                                    Some r' \<Rightarrow> Some (Node l' y (d \<or> (x=y)) r')
62f6ba39b3d4 modernized specifications;
wenzelm
parents: 35408
diff changeset
   108
                                  | None \<Rightarrow> Some (Node l' y (d \<or> (x=y)) r))
62f6ba39b3d4 modernized specifications;
wenzelm
parents: 35408
diff changeset
   109
                               | None \<Rightarrow>
45358
4849133d7a78 tuned document;
wenzelm
parents: 45355
diff changeset
   110
                                  (case delete x r of 
38838
62f6ba39b3d4 modernized specifications;
wenzelm
parents: 35408
diff changeset
   111
                                     Some r' \<Rightarrow> Some (Node l y (d \<or> (x=y)) r')
62f6ba39b3d4 modernized specifications;
wenzelm
parents: 35408
diff changeset
   112
                                   | None \<Rightarrow> if x=y \<and> \<not>d then Some (Node l y True r)
62f6ba39b3d4 modernized specifications;
wenzelm
parents: 35408
diff changeset
   113
                                             else None))"
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   114
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   115
45358
4849133d7a78 tuned document;
wenzelm
parents: 45355
diff changeset
   116
lemma delete_Some_set_of: "delete x t = Some t' \<Longrightarrow> set_of t' \<subseteq> set_of t"
4849133d7a78 tuned document;
wenzelm
parents: 45355
diff changeset
   117
proof (induct t arbitrary: t')
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   118
  case Tip thus ?case by simp
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   119
next
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   120
  case (Node l y d r)
25364
7f012f56efa3 tuned proofs -- avoid implicit prems;
wenzelm
parents: 25174
diff changeset
   121
  have del: "delete x (Node l y d r) = Some t'" by fact
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   122
  show ?case
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   123
  proof (cases "delete x l")
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   124
    case (Some l')
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   125
    note x_l_Some = this
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   126
    with Node.hyps
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   127
    have l'_l: "set_of l' \<subseteq> set_of l"
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   128
      by simp
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   129
    show ?thesis
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   130
    proof (cases "delete x r")
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   131
      case (Some r')
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   132
      with Node.hyps
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   133
      have "set_of r' \<subseteq> set_of r"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   134
        by simp
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   135
      with l'_l Some x_l_Some del
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   136
      show ?thesis
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   137
        by (auto split: split_if_asm)
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   138
    next
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   139
      case None
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   140
      with l'_l Some x_l_Some del
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   141
      show ?thesis
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42287
diff changeset
   142
        by (fastforce split: split_if_asm)
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   143
    qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   144
  next
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   145
    case None
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   146
    note x_l_None = this
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   147
    show ?thesis
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   148
    proof (cases "delete x r")
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   149
      case (Some r')
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   150
      with Node.hyps
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   151
      have "set_of r' \<subseteq> set_of r"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   152
        by simp
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   153
      with Some x_l_None del
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   154
      show ?thesis
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42287
diff changeset
   155
        by (fastforce split: split_if_asm)
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   156
    next
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   157
      case None
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   158
      with x_l_None del
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   159
      show ?thesis
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42287
diff changeset
   160
        by (fastforce split: split_if_asm)
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   161
    qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   162
  qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   163
qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   164
45355
c0704e988526 misc tuning and modernization;
wenzelm
parents: 44890
diff changeset
   165
lemma delete_Some_all_distinct:
45358
4849133d7a78 tuned document;
wenzelm
parents: 45355
diff changeset
   166
  "delete x t = Some t' \<Longrightarrow> all_distinct t \<Longrightarrow> all_distinct t'"
4849133d7a78 tuned document;
wenzelm
parents: 45355
diff changeset
   167
proof (induct t arbitrary: t')
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   168
  case Tip thus ?case by simp
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   169
next
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   170
  case (Node l y d r)
25364
7f012f56efa3 tuned proofs -- avoid implicit prems;
wenzelm
parents: 25174
diff changeset
   171
  have del: "delete x (Node l y d r) = Some t'" by fact
7f012f56efa3 tuned proofs -- avoid implicit prems;
wenzelm
parents: 25174
diff changeset
   172
  have "all_distinct (Node l y d r)" by fact
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   173
  then obtain
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   174
    dist_l: "all_distinct l" and
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   175
    dist_r: "all_distinct r" and
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   176
    d: "d \<or> (y \<notin> set_of l \<and> y \<notin> set_of r)" and
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   177
    dist_l_r: "set_of l \<inter> set_of r = {}"
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   178
    by auto
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   179
  show ?case
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   180
  proof (cases "delete x l")
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   181
    case (Some l')
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   182
    note x_l_Some = this
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   183
    from Node.hyps (1) [OF Some dist_l]
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   184
    have dist_l': "all_distinct l'"
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   185
      by simp
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   186
    from delete_Some_set_of [OF x_l_Some]
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   187
    have l'_l: "set_of l' \<subseteq> set_of l".
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   188
    show ?thesis
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   189
    proof (cases "delete x r")
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   190
      case (Some r')
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   191
      from Node.hyps (2) [OF Some dist_r]
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   192
      have dist_r': "all_distinct r'"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   193
        by simp
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   194
      from delete_Some_set_of [OF Some]
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   195
      have "set_of r' \<subseteq> set_of r".
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   196
      
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   197
      with dist_l' dist_r' l'_l Some x_l_Some del d dist_l_r
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   198
      show ?thesis
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42287
diff changeset
   199
        by fastforce
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   200
    next
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   201
      case None
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   202
      with l'_l dist_l'  x_l_Some del d dist_l_r dist_r
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   203
      show ?thesis
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42287
diff changeset
   204
        by fastforce
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   205
    qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   206
  next
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   207
    case None
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   208
    note x_l_None = this
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   209
    show ?thesis
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   210
    proof (cases "delete x r")
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   211
      case (Some r')
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   212
      with Node.hyps (2) [OF Some dist_r]
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   213
      have dist_r': "all_distinct r'"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   214
        by simp
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   215
      from delete_Some_set_of [OF Some]
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   216
      have "set_of r' \<subseteq> set_of r".
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   217
      with Some dist_r' x_l_None del dist_l d dist_l_r
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   218
      show ?thesis
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42287
diff changeset
   219
        by fastforce
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   220
    next
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   221
      case None
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   222
      with x_l_None del dist_l dist_r d dist_l_r
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   223
      show ?thesis
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42287
diff changeset
   224
        by (fastforce split: split_if_asm)
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   225
    qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   226
  qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   227
qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   228
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   229
lemma delete_None_set_of_conv: "delete x t = None = (x \<notin> set_of t)"
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   230
proof (induct t)
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   231
  case Tip thus ?case by simp
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   232
next
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   233
  case (Node l y d r)
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   234
  thus ?case
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   235
    by (auto split: option.splits)
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   236
qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   237
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   238
lemma delete_Some_x_set_of:
45358
4849133d7a78 tuned document;
wenzelm
parents: 45355
diff changeset
   239
  "delete x t = Some t' \<Longrightarrow> x \<in> set_of t \<and> x \<notin> set_of t'"
4849133d7a78 tuned document;
wenzelm
parents: 45355
diff changeset
   240
proof (induct t arbitrary: t')
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   241
  case Tip thus ?case by simp
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   242
next
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   243
  case (Node l y d r)
25364
7f012f56efa3 tuned proofs -- avoid implicit prems;
wenzelm
parents: 25174
diff changeset
   244
  have del: "delete x (Node l y d r) = Some t'" by fact
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   245
  show ?case
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   246
  proof (cases "delete x l")
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   247
    case (Some l')
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   248
    note x_l_Some = this
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   249
    from Node.hyps (1) [OF Some]
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   250
    obtain x_l: "x \<in> set_of l" "x \<notin> set_of l'"
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   251
      by simp
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   252
    show ?thesis
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   253
    proof (cases "delete x r")
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   254
      case (Some r')
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   255
      from Node.hyps (2) [OF Some]
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   256
      obtain x_r: "x \<in> set_of r" "x \<notin> set_of r'"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   257
        by simp
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   258
      from x_r x_l Some x_l_Some del 
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   259
      show ?thesis
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   260
        by (clarsimp split: split_if_asm)
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   261
    next
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   262
      case None
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   263
      then have "x \<notin> set_of r"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   264
        by (simp add: delete_None_set_of_conv)
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   265
      with x_l None x_l_Some del
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   266
      show ?thesis
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   267
        by (clarsimp split: split_if_asm)
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   268
    qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   269
  next
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   270
    case None
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   271
    note x_l_None = this
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   272
    then have x_notin_l: "x \<notin> set_of l"
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   273
      by (simp add: delete_None_set_of_conv)
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   274
    show ?thesis
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   275
    proof (cases "delete x r")
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   276
      case (Some r')
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   277
      from Node.hyps (2) [OF Some]
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   278
      obtain x_r: "x \<in> set_of r" "x \<notin> set_of r'"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   279
        by simp
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   280
      from x_r x_notin_l Some x_l_None del 
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   281
      show ?thesis
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   282
        by (clarsimp split: split_if_asm)
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   283
    next
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   284
      case None
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   285
      then have "x \<notin> set_of r"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   286
        by (simp add: delete_None_set_of_conv)
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   287
      with None x_l_None x_notin_l del
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   288
      show ?thesis
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   289
        by (clarsimp split: split_if_asm)
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   290
    qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   291
  qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   292
qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   293
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   294
38838
62f6ba39b3d4 modernized specifications;
wenzelm
parents: 35408
diff changeset
   295
primrec subtract :: "'a tree \<Rightarrow> 'a tree \<Rightarrow> 'a tree option"
62f6ba39b3d4 modernized specifications;
wenzelm
parents: 35408
diff changeset
   296
where
62f6ba39b3d4 modernized specifications;
wenzelm
parents: 35408
diff changeset
   297
  "subtract Tip t = Some t"
62f6ba39b3d4 modernized specifications;
wenzelm
parents: 35408
diff changeset
   298
| "subtract (Node l x b r) t =
62f6ba39b3d4 modernized specifications;
wenzelm
parents: 35408
diff changeset
   299
     (case delete x t of
62f6ba39b3d4 modernized specifications;
wenzelm
parents: 35408
diff changeset
   300
        Some t' \<Rightarrow> (case subtract l t' of 
62f6ba39b3d4 modernized specifications;
wenzelm
parents: 35408
diff changeset
   301
                     Some t'' \<Rightarrow> subtract r t''
62f6ba39b3d4 modernized specifications;
wenzelm
parents: 35408
diff changeset
   302
                    | None \<Rightarrow> None)
62f6ba39b3d4 modernized specifications;
wenzelm
parents: 35408
diff changeset
   303
       | None \<Rightarrow> None)"
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   304
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   305
lemma subtract_Some_set_of_res: 
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   306
  "subtract t\<^sub>1 t\<^sub>2 = Some t \<Longrightarrow> set_of t \<subseteq> set_of t\<^sub>2"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   307
proof (induct t\<^sub>1 arbitrary: t\<^sub>2 t)
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   308
  case Tip thus ?case by simp
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   309
next
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   310
  case (Node l x b r)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   311
  have sub: "subtract (Node l x b r) t\<^sub>2 = Some t" by fact
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   312
  show ?case
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   313
  proof (cases "delete x t\<^sub>2")
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   314
    case (Some t\<^sub>2')
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   315
    note del_x_Some = this
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   316
    from delete_Some_set_of [OF Some] 
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   317
    have t2'_t2: "set_of t\<^sub>2' \<subseteq> set_of t\<^sub>2" .
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   318
    show ?thesis
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   319
    proof (cases "subtract l t\<^sub>2'")
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   320
      case (Some t\<^sub>2'')
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   321
      note sub_l_Some = this
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   322
      from Node.hyps (1) [OF Some] 
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   323
      have t2''_t2': "set_of t\<^sub>2'' \<subseteq> set_of t\<^sub>2'" .
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   324
      show ?thesis
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   325
      proof (cases "subtract r t\<^sub>2''")
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   326
        case (Some t\<^sub>2''')
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   327
        from Node.hyps (2) [OF Some ] 
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   328
        have "set_of t\<^sub>2''' \<subseteq> set_of t\<^sub>2''" .
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   329
        with Some sub_l_Some del_x_Some sub t2''_t2' t2'_t2
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   330
        show ?thesis
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   331
          by simp
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   332
      next
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   333
        case None
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   334
        with del_x_Some sub_l_Some sub
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   335
        show ?thesis
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   336
          by simp
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   337
      qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   338
    next
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   339
      case None
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   340
      with del_x_Some sub 
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   341
      show ?thesis
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   342
        by simp
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   343
    qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   344
  next
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   345
    case None
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   346
    with sub show ?thesis by simp
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   347
  qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   348
qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   349
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   350
lemma subtract_Some_set_of: 
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   351
  "subtract t\<^sub>1 t\<^sub>2 = Some t \<Longrightarrow> set_of t\<^sub>1 \<subseteq> set_of t\<^sub>2"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   352
proof (induct t\<^sub>1 arbitrary: t\<^sub>2 t)
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   353
  case Tip thus ?case by simp
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   354
next
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   355
  case (Node l x d r)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   356
  have sub: "subtract (Node l x d r) t\<^sub>2 = Some t" by fact
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   357
  show ?case
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   358
  proof (cases "delete x t\<^sub>2")
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   359
    case (Some t\<^sub>2')
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   360
    note del_x_Some = this
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   361
    from delete_Some_set_of [OF Some] 
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   362
    have t2'_t2: "set_of t\<^sub>2' \<subseteq> set_of t\<^sub>2" .
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   363
    from delete_None_set_of_conv [of x t\<^sub>2] Some
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   364
    have x_t2: "x \<in> set_of t\<^sub>2"
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   365
      by simp
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   366
    show ?thesis
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   367
    proof (cases "subtract l t\<^sub>2'")
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   368
      case (Some t\<^sub>2'')
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   369
      note sub_l_Some = this
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   370
      from Node.hyps (1) [OF Some] 
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   371
      have l_t2': "set_of l \<subseteq> set_of t\<^sub>2'" .
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   372
      from subtract_Some_set_of_res [OF Some]
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   373
      have t2''_t2': "set_of t\<^sub>2'' \<subseteq> set_of t\<^sub>2'" .
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   374
      show ?thesis
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   375
      proof (cases "subtract r t\<^sub>2''")
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   376
        case (Some t\<^sub>2''')
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   377
        from Node.hyps (2) [OF Some ] 
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   378
        have r_t\<^sub>2'': "set_of r \<subseteq> set_of t\<^sub>2''" .
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   379
        from Some sub_l_Some del_x_Some sub r_t\<^sub>2'' l_t2' t2'_t2 t2''_t2' x_t2
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   380
        show ?thesis
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   381
          by auto
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   382
      next
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   383
        case None
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   384
        with del_x_Some sub_l_Some sub
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   385
        show ?thesis
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   386
          by simp
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   387
      qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   388
    next
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   389
      case None
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   390
      with del_x_Some sub 
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   391
      show ?thesis
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   392
        by simp
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   393
    qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   394
  next
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   395
    case None
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   396
    with sub show ?thesis by simp
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   397
  qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   398
qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   399
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   400
lemma subtract_Some_all_distinct_res: 
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   401
  "subtract t\<^sub>1 t\<^sub>2 = Some t \<Longrightarrow> all_distinct t\<^sub>2 \<Longrightarrow> all_distinct t"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   402
proof (induct t\<^sub>1 arbitrary: t\<^sub>2 t)
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   403
  case Tip thus ?case by simp
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   404
next
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   405
  case (Node l x d r)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   406
  have sub: "subtract (Node l x d r) t\<^sub>2 = Some t" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   407
  have dist_t2: "all_distinct t\<^sub>2" by fact
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   408
  show ?case
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   409
  proof (cases "delete x t\<^sub>2")
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   410
    case (Some t\<^sub>2')
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   411
    note del_x_Some = this
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   412
    from delete_Some_all_distinct [OF Some dist_t2] 
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   413
    have dist_t2': "all_distinct t\<^sub>2'" .
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   414
    show ?thesis
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   415
    proof (cases "subtract l t\<^sub>2'")
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   416
      case (Some t\<^sub>2'')
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   417
      note sub_l_Some = this
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   418
      from Node.hyps (1) [OF Some dist_t2'] 
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   419
      have dist_t2'': "all_distinct t\<^sub>2''" .
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   420
      show ?thesis
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   421
      proof (cases "subtract r t\<^sub>2''")
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   422
        case (Some t\<^sub>2''')
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   423
        from Node.hyps (2) [OF Some dist_t2''] 
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   424
        have dist_t2''': "all_distinct t\<^sub>2'''" .
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   425
        from Some sub_l_Some del_x_Some sub 
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   426
             dist_t2'''
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   427
        show ?thesis
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   428
          by simp
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   429
      next
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   430
        case None
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   431
        with del_x_Some sub_l_Some sub
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   432
        show ?thesis
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   433
          by simp
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   434
      qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   435
    next
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   436
      case None
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   437
      with del_x_Some sub 
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   438
      show ?thesis
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   439
        by simp
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   440
    qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   441
  next
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   442
    case None
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   443
    with sub show ?thesis by simp
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   444
  qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   445
qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   446
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   447
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   448
lemma subtract_Some_dist_res: 
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   449
  "subtract t\<^sub>1 t\<^sub>2 = Some t \<Longrightarrow> set_of t\<^sub>1 \<inter> set_of t = {}"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   450
proof (induct t\<^sub>1 arbitrary: t\<^sub>2 t)
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   451
  case Tip thus ?case by simp
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   452
next
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   453
  case (Node l x d r)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   454
  have sub: "subtract (Node l x d r) t\<^sub>2 = Some t" by fact
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   455
  show ?case
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   456
  proof (cases "delete x t\<^sub>2")
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   457
    case (Some t\<^sub>2')
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   458
    note del_x_Some = this
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   459
    from delete_Some_x_set_of [OF Some]
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   460
    obtain x_t2: "x \<in> set_of t\<^sub>2" and x_not_t2': "x \<notin> set_of t\<^sub>2'"
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   461
      by simp
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   462
    from delete_Some_set_of [OF Some]
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   463
    have t2'_t2: "set_of t\<^sub>2' \<subseteq> set_of t\<^sub>2" .
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   464
    show ?thesis
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   465
    proof (cases "subtract l t\<^sub>2'")
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   466
      case (Some t\<^sub>2'')
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   467
      note sub_l_Some = this
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   468
      from Node.hyps (1) [OF Some ] 
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   469
      have dist_l_t2'': "set_of l \<inter> set_of t\<^sub>2'' = {}".
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   470
      from subtract_Some_set_of_res [OF Some]
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   471
      have t2''_t2': "set_of t\<^sub>2'' \<subseteq> set_of t\<^sub>2'" .
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   472
      show ?thesis
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   473
      proof (cases "subtract r t\<^sub>2''")
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   474
        case (Some t\<^sub>2''')
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   475
        from Node.hyps (2) [OF Some] 
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   476
        have dist_r_t2''': "set_of r \<inter> set_of t\<^sub>2''' = {}" .
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   477
        from subtract_Some_set_of_res [OF Some]
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   478
        have t2'''_t2'': "set_of t\<^sub>2''' \<subseteq> set_of t\<^sub>2''".
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   479
        
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   480
        from Some sub_l_Some del_x_Some sub t2'''_t2'' dist_l_t2'' dist_r_t2'''
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   481
             t2''_t2' t2'_t2 x_not_t2'
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   482
        show ?thesis
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   483
          by auto
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   484
      next
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   485
        case None
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   486
        with del_x_Some sub_l_Some sub
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   487
        show ?thesis
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   488
          by simp
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   489
      qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   490
    next
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   491
      case None
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   492
      with del_x_Some sub 
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   493
      show ?thesis
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   494
        by simp
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   495
    qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   496
  next
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   497
    case None
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   498
    with sub show ?thesis by simp
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   499
  qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   500
qed
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   501
        
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   502
lemma subtract_Some_all_distinct:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   503
  "subtract t\<^sub>1 t\<^sub>2 = Some t \<Longrightarrow> all_distinct t\<^sub>2 \<Longrightarrow> all_distinct t\<^sub>1"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   504
proof (induct t\<^sub>1 arbitrary: t\<^sub>2 t)
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   505
  case Tip thus ?case by simp
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   506
next
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   507
  case (Node l x d r)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   508
  have sub: "subtract (Node l x d r) t\<^sub>2 = Some t" by fact
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   509
  have dist_t2: "all_distinct t\<^sub>2" by fact
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   510
  show ?case
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   511
  proof (cases "delete x t\<^sub>2")
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   512
    case (Some t\<^sub>2')
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   513
    note del_x_Some = this
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   514
    from delete_Some_all_distinct [OF Some dist_t2 ] 
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   515
    have dist_t2': "all_distinct t\<^sub>2'" .
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   516
    from delete_Some_set_of [OF Some]
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   517
    have t2'_t2: "set_of t\<^sub>2' \<subseteq> set_of t\<^sub>2" .
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   518
    from delete_Some_x_set_of [OF Some]
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   519
    obtain x_t2: "x \<in> set_of t\<^sub>2" and x_not_t2': "x \<notin> set_of t\<^sub>2'"
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   520
      by simp
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   521
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   522
    show ?thesis
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   523
    proof (cases "subtract l t\<^sub>2'")
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   524
      case (Some t\<^sub>2'')
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   525
      note sub_l_Some = this
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   526
      from Node.hyps (1) [OF Some dist_t2' ] 
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   527
      have dist_l: "all_distinct l" .
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   528
      from subtract_Some_all_distinct_res [OF Some dist_t2'] 
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   529
      have dist_t2'': "all_distinct t\<^sub>2''" .
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   530
      from subtract_Some_set_of [OF Some]
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   531
      have l_t2': "set_of l \<subseteq> set_of t\<^sub>2'" .
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   532
      from subtract_Some_set_of_res [OF Some]
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   533
      have t2''_t2': "set_of t\<^sub>2'' \<subseteq> set_of t\<^sub>2'" .
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   534
      from subtract_Some_dist_res [OF Some]
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   535
      have dist_l_t2'': "set_of l \<inter> set_of t\<^sub>2'' = {}".
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   536
      show ?thesis
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   537
      proof (cases "subtract r t\<^sub>2''")
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   538
        case (Some t\<^sub>2''')
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   539
        from Node.hyps (2) [OF Some dist_t2''] 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   540
        have dist_r: "all_distinct r" .
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   541
        from subtract_Some_set_of [OF Some]
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   542
        have r_t2'': "set_of r \<subseteq> set_of t\<^sub>2''" .
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   543
        from subtract_Some_dist_res [OF Some]
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 48891
diff changeset
   544
        have dist_r_t2''': "set_of r \<inter> set_of t\<^sub>2''' = {}".
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   545
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   546
        from dist_l dist_r Some sub_l_Some del_x_Some r_t2'' l_t2' x_t2 x_not_t2' 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   547
             t2''_t2' dist_l_t2'' dist_r_t2'''
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   548
        show ?thesis
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   549
          by auto
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   550
      next
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   551
        case None
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   552
        with del_x_Some sub_l_Some sub
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   553
        show ?thesis
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   554
          by simp
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   555
      qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   556
    next
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   557
      case None
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   558
      with del_x_Some sub 
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   559
      show ?thesis
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
   560
        by simp
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   561
    qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   562
  next
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   563
    case None
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   564
    with sub show ?thesis by simp
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   565
  qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   566
qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   567
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   568
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   569
lemma delete_left:
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   570
  assumes dist: "all_distinct (Node l y d r)" 
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   571
  assumes del_l: "delete x l = Some l'"
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   572
  shows "delete x (Node l y d r) = Some (Node l' y d r)"
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   573
proof -
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   574
  from delete_Some_x_set_of [OF del_l]
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 53015
diff changeset
   575
  obtain x: "x \<in> set_of l"
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   576
    by simp
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 53015
diff changeset
   577
  with dist 
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   578
  have "delete x r = None"
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   579
    by (cases "delete x r") (auto dest:delete_Some_x_set_of)
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   580
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 53015
diff changeset
   581
  with x 
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   582
  show ?thesis
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   583
    using del_l dist
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   584
    by (auto split: option.splits)
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   585
qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   586
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   587
lemma delete_right:
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   588
  assumes dist: "all_distinct (Node l y d r)" 
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   589
  assumes del_r: "delete x r = Some r'"
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   590
  shows "delete x (Node l y d r) = Some (Node l y d r')"
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   591
proof -
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   592
  from delete_Some_x_set_of [OF del_r]
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 53015
diff changeset
   593
  obtain x: "x \<in> set_of r"
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   594
    by simp
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 53015
diff changeset
   595
  with dist 
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   596
  have "delete x l = None"
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   597
    by (cases "delete x l") (auto dest:delete_Some_x_set_of)
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   598
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 53015
diff changeset
   599
  with x 
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   600
  show ?thesis
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   601
    using del_r dist
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   602
    by (auto split: option.splits)
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   603
qed
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   604
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   605
lemma delete_root: 
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   606
  assumes dist: "all_distinct (Node l x False r)" 
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   607
  shows "delete x (Node l x False r) = Some (Node l x True r)"
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   608
proof -
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   609
  from dist have "delete x r = None"
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   610
    by (cases "delete x r") (auto dest:delete_Some_x_set_of)
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   611
  moreover
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   612
  from dist have "delete x l = None"
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   613
    by (cases "delete x l") (auto dest:delete_Some_x_set_of)
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   614
  ultimately show ?thesis
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   615
    using dist
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   616
       by (auto split: option.splits)
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   617
qed               
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   618
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   619
lemma subtract_Node:
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   620
 assumes del: "delete x t = Some t'"                                
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   621
 assumes sub_l: "subtract l t' = Some t''"
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   622
 assumes sub_r: "subtract r t'' = Some t'''"
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   623
 shows "subtract (Node l x False r) t = Some t'''"
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   624
using del sub_l sub_r
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   625
by simp
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   626
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   627
lemma subtract_Tip: "subtract Tip t = Some t"
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   628
  by simp
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   629
 
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   630
text {* Now we have all the theorems in place that are needed for the
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   631
certificate generating ML functions. *}
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   632
48891
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 45358
diff changeset
   633
ML_file "distinct_tree_prover.ML"
25171
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   634
4a9c25bffc9b added Statespace library
schirmer
parents:
diff changeset
   635
end