doc-src/ZF/ZF_examples.thy
author Walther Neuper <neuper@ist.tugraz.at>
Thu, 12 Aug 2010 15:03:34 +0200
branchisac-from-Isabelle2009-2
changeset 37913 20e3616b2d9c
parent 35413 d8d7d1b785af
child 47347 2289a3869c88
permissions -rw-r--r--
prepare reactivation of isac-update-Isa09-2
paulson@14152
     1
header{*Examples of Reasoning in ZF Set Theory*}
paulson@14152
     2
haftmann@16417
     3
theory ZF_examples imports Main_ZFC begin
paulson@14152
     4
paulson@14152
     5
subsection {* Binary Trees *}
paulson@14152
     6
paulson@14152
     7
consts
paulson@14152
     8
  bt :: "i => i"
paulson@14152
     9
paulson@14152
    10
datatype "bt(A)" =
paulson@14152
    11
  Lf | Br ("a \<in> A", "t1 \<in> bt(A)", "t2 \<in> bt(A)")
paulson@14152
    12
paulson@14152
    13
declare bt.intros [simp]
paulson@14152
    14
paulson@14152
    15
text{*Induction via tactic emulation*}
paulson@14152
    16
lemma Br_neq_left [rule_format]: "l \<in> bt(A) ==> \<forall>x r. Br(x, l, r) \<noteq> l"
paulson@14152
    17
  --{* @{subgoals[display,indent=0,margin=65]} *}
paulson@14152
    18
  apply (induct_tac l)
paulson@14152
    19
  --{* @{subgoals[display,indent=0,margin=65]} *}
paulson@14152
    20
  apply auto
paulson@14152
    21
  done
paulson@14152
    22
paulson@14152
    23
(*
paulson@14152
    24
  apply (Inductive.case_tac l)
paulson@14152
    25
  apply (tactic {*exhaust_tac "l" 1*})
paulson@14152
    26
*)
paulson@14152
    27
paulson@14152
    28
text{*The new induction method, which I don't understand*}
paulson@14152
    29
lemma Br_neq_left': "l \<in> bt(A) ==> (!!x r. Br(x, l, r) \<noteq> l)"
paulson@14152
    30
  --{* @{subgoals[display,indent=0,margin=65]} *}
paulson@14152
    31
  apply (induct set: bt)
paulson@14152
    32
  --{* @{subgoals[display,indent=0,margin=65]} *}
paulson@14152
    33
  apply auto
paulson@14152
    34
  done
paulson@14152
    35
paulson@14159
    36
lemma Br_iff: "Br(a,l,r) = Br(a',l',r') <-> a=a' & l=l' & r=r'"
paulson@14152
    37
  -- "Proving a freeness theorem."
paulson@14152
    38
  by (blast elim!: bt.free_elims)
paulson@14152
    39
paulson@14159
    40
inductive_cases Br_in_bt: "Br(a,l,r) \<in> bt(A)"
paulson@14152
    41
  -- "An elimination rule, for type-checking."
paulson@14152
    42
paulson@14152
    43
text {*
paulson@14159
    44
@{thm[display] Br_in_bt[no_vars]}
paulson@14152
    45
*};
paulson@14152
    46
paulson@14159
    47
subsection{*Primitive recursion*}
paulson@14159
    48
paulson@14159
    49
consts  n_nodes :: "i => i"
paulson@14159
    50
primrec
paulson@14159
    51
  "n_nodes(Lf) = 0"
paulson@14159
    52
  "n_nodes(Br(a,l,r)) = succ(n_nodes(l) #+ n_nodes(r))"
paulson@14159
    53
paulson@14159
    54
lemma n_nodes_type [simp]: "t \<in> bt(A) ==> n_nodes(t) \<in> nat"
paulson@14159
    55
  by (induct_tac t, auto) 
paulson@14159
    56
paulson@14159
    57
consts  n_nodes_aux :: "i => i"
paulson@14159
    58
primrec
paulson@14159
    59
  "n_nodes_aux(Lf) = (\<lambda>k \<in> nat. k)"
paulson@14159
    60
  "n_nodes_aux(Br(a,l,r)) = 
paulson@14159
    61
      (\<lambda>k \<in> nat. n_nodes_aux(r) `  (n_nodes_aux(l) ` succ(k)))"
paulson@14159
    62
paulson@14159
    63
lemma n_nodes_aux_eq [rule_format]:
paulson@14159
    64
     "t \<in> bt(A) ==> \<forall>k \<in> nat. n_nodes_aux(t)`k = n_nodes(t) #+ k"
paulson@14159
    65
  by (induct_tac t, simp_all) 
paulson@14159
    66
haftmann@35413
    67
definition n_nodes_tail :: "i => i" where
paulson@14159
    68
   "n_nodes_tail(t) == n_nodes_aux(t) ` 0"
paulson@14159
    69
paulson@14159
    70
lemma "t \<in> bt(A) ==> n_nodes_tail(t) = n_nodes(t)"
paulson@14159
    71
 by (simp add: n_nodes_tail_def n_nodes_aux_eq) 
paulson@14159
    72
paulson@14159
    73
paulson@14159
    74
subsection {*Inductive definitions*}
paulson@14159
    75
paulson@14159
    76
consts  Fin       :: "i=>i"
paulson@14159
    77
inductive
paulson@14159
    78
  domains   "Fin(A)" \<subseteq> "Pow(A)"
paulson@14159
    79
  intros
paulson@14159
    80
    emptyI:  "0 \<in> Fin(A)"
paulson@14159
    81
    consI:   "[| a \<in> A;  b \<in> Fin(A) |] ==> cons(a,b) \<in> Fin(A)"
paulson@14159
    82
  type_intros  empty_subsetI cons_subsetI PowI
paulson@14159
    83
  type_elims   PowD [THEN revcut_rl]
paulson@14159
    84
paulson@14159
    85
paulson@14159
    86
consts  acc :: "i => i"
paulson@14159
    87
inductive
paulson@14159
    88
  domains "acc(r)" \<subseteq> "field(r)"
paulson@14159
    89
  intros
paulson@14159
    90
    vimage:  "[| r-``{a}: Pow(acc(r)); a \<in> field(r) |] ==> a \<in> acc(r)"
paulson@14159
    91
  monos      Pow_mono
paulson@14159
    92
paulson@14159
    93
paulson@14159
    94
consts
paulson@14159
    95
  llist  :: "i=>i";
paulson@14159
    96
paulson@14159
    97
codatatype
paulson@14159
    98
  "llist(A)" = LNil | LCons ("a \<in> A", "l \<in> llist(A)")
paulson@14159
    99
paulson@14159
   100
paulson@14159
   101
(*Coinductive definition of equality*)
paulson@14159
   102
consts
paulson@14159
   103
  lleq :: "i=>i"
paulson@14159
   104
paulson@14159
   105
(*Previously used <*> in the domain and variant pairs as elements.  But
paulson@14159
   106
  standard pairs work just as well.  To use variant pairs, must change prefix
paulson@14159
   107
  a q/Q to the Sigma, Pair and converse rules.*)
paulson@14159
   108
coinductive
paulson@14159
   109
  domains "lleq(A)" \<subseteq> "llist(A) * llist(A)"
paulson@14159
   110
  intros
paulson@14159
   111
    LNil:  "<LNil, LNil> \<in> lleq(A)"
paulson@14159
   112
    LCons: "[| a \<in> A; <l,l'> \<in> lleq(A) |] 
paulson@14159
   113
            ==> <LCons(a,l), LCons(a,l')> \<in> lleq(A)"
paulson@14159
   114
  type_intros  llist.intros
paulson@14159
   115
paulson@14159
   116
paulson@14152
   117
subsection{*Powerset example*}
paulson@14152
   118
paulson@14159
   119
lemma Pow_mono: "A\<subseteq>B  ==>  Pow(A) \<subseteq> Pow(B)"
paulson@14152
   120
apply (rule subsetI)
paulson@14152
   121
apply (rule PowI)
paulson@14152
   122
apply (drule PowD)
paulson@14152
   123
apply (erule subset_trans, assumption)
paulson@14152
   124
done
paulson@14152
   125
paulson@14152
   126
lemma "Pow(A Int B) = Pow(A) Int Pow(B)"
paulson@14152
   127
  --{* @{subgoals[display,indent=0,margin=65]} *}
paulson@14152
   128
apply (rule equalityI)
paulson@14152
   129
  --{* @{subgoals[display,indent=0,margin=65]} *}
paulson@14152
   130
apply (rule Int_greatest)
paulson@14152
   131
  --{* @{subgoals[display,indent=0,margin=65]} *}
paulson@14152
   132
apply (rule Int_lower1 [THEN Pow_mono])
paulson@14152
   133
  --{* @{subgoals[display,indent=0,margin=65]} *}
paulson@14152
   134
apply (rule Int_lower2 [THEN Pow_mono])
paulson@14152
   135
  --{* @{subgoals[display,indent=0,margin=65]} *}
paulson@14152
   136
apply (rule subsetI)
paulson@14152
   137
  --{* @{subgoals[display,indent=0,margin=65]} *}
paulson@14152
   138
apply (erule IntE)
paulson@14152
   139
  --{* @{subgoals[display,indent=0,margin=65]} *}
paulson@14152
   140
apply (rule PowI)
paulson@14152
   141
  --{* @{subgoals[display,indent=0,margin=65]} *}
paulson@14152
   142
apply (drule PowD)+
paulson@14152
   143
  --{* @{subgoals[display,indent=0,margin=65]} *}
paulson@14159
   144
apply (rule Int_greatest)
paulson@14159
   145
  --{* @{subgoals[display,indent=0,margin=65]} *}
paulson@14159
   146
apply (assumption+)
paulson@14152
   147
done
paulson@14152
   148
paulson@14152
   149
text{*Trying again from the beginning in order to use @{text blast}*}
paulson@14152
   150
lemma "Pow(A Int B) = Pow(A) Int Pow(B)"
paulson@14152
   151
by blast
paulson@14152
   152
paulson@14152
   153
paulson@14159
   154
lemma "C\<subseteq>D ==> Union(C) \<subseteq> Union(D)"
paulson@14152
   155
  --{* @{subgoals[display,indent=0,margin=65]} *}
paulson@14152
   156
apply (rule subsetI)
paulson@14152
   157
  --{* @{subgoals[display,indent=0,margin=65]} *}
paulson@14152
   158
apply (erule UnionE)
paulson@14152
   159
  --{* @{subgoals[display,indent=0,margin=65]} *}
paulson@14152
   160
apply (rule UnionI)
paulson@14152
   161
  --{* @{subgoals[display,indent=0,margin=65]} *}
paulson@14159
   162
apply (erule subsetD)
paulson@14159
   163
  --{* @{subgoals[display,indent=0,margin=65]} *}
paulson@14159
   164
apply assumption 
paulson@14159
   165
  --{* @{subgoals[display,indent=0,margin=65]} *}
paulson@14159
   166
apply assumption 
paulson@14152
   167
done
paulson@14152
   168
paulson@14159
   169
text{*A more abstract version of the same proof*}
paulson@14152
   170
paulson@14159
   171
lemma "C\<subseteq>D ==> Union(C) \<subseteq> Union(D)"
paulson@14152
   172
  --{* @{subgoals[display,indent=0,margin=65]} *}
paulson@14152
   173
apply (rule Union_least)
paulson@14152
   174
  --{* @{subgoals[display,indent=0,margin=65]} *}
paulson@14152
   175
apply (rule Union_upper)
paulson@14152
   176
  --{* @{subgoals[display,indent=0,margin=65]} *}
paulson@14152
   177
apply (erule subsetD, assumption)
paulson@14152
   178
done
paulson@14152
   179
paulson@14152
   180
paulson@14159
   181
lemma "[| a \<in> A;  f \<in> A->B;  g \<in> C->D;  A \<inter> C = 0 |] ==> (f \<union> g)`a = f`a"
paulson@14152
   182
  --{* @{subgoals[display,indent=0,margin=65]} *}
paulson@14152
   183
apply (rule apply_equality)
paulson@14152
   184
  --{* @{subgoals[display,indent=0,margin=65]} *}
paulson@14152
   185
apply (rule UnI1)
paulson@14152
   186
  --{* @{subgoals[display,indent=0,margin=65]} *}
paulson@14159
   187
apply (rule apply_Pair)
paulson@14152
   188
  --{* @{subgoals[display,indent=0,margin=65]} *}
paulson@14159
   189
apply assumption 
paulson@14159
   190
  --{* @{subgoals[display,indent=0,margin=65]} *}
paulson@14159
   191
apply assumption 
paulson@14159
   192
  --{* @{subgoals[display,indent=0,margin=65]} *}
paulson@14159
   193
apply (rule fun_disjoint_Un)
paulson@14159
   194
  --{* @{subgoals[display,indent=0,margin=65]} *}
paulson@14159
   195
apply assumption 
paulson@14159
   196
  --{* @{subgoals[display,indent=0,margin=65]} *}
paulson@14159
   197
apply assumption 
paulson@14159
   198
  --{* @{subgoals[display,indent=0,margin=65]} *}
paulson@14159
   199
apply assumption 
paulson@14152
   200
done
paulson@14152
   201
paulson@14152
   202
end