src/ZF/Resid/Redex.thy
author paulson
Tue, 06 Mar 2012 17:01:37 +0000
changeset 46823 57bf0cecb366
parent 32960 69916a850301
child 61398 6794de675568
permissions -rw-r--r--
More mathematical symbols for ZF examples
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9284
85a5355faa91 removal of (harmless) circular definitions
paulson
parents: 6112
diff changeset
     1
(*  Title:      ZF/Resid/Redex.thy
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 22808
diff changeset
     2
    Author:     Ole Rasmussen, University of Cambridge
1048
5ba0314f8214 New example by Ole Rasmussen
lcp
parents:
diff changeset
     3
*)
5ba0314f8214 New example by Ole Rasmussen
lcp
parents:
diff changeset
     4
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 13612
diff changeset
     5
theory Redex imports Main begin
1048
5ba0314f8214 New example by Ole Rasmussen
lcp
parents:
diff changeset
     6
consts
1401
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 1155
diff changeset
     7
  redexes     :: i
1048
5ba0314f8214 New example by Ole Rasmussen
lcp
parents:
diff changeset
     8
5ba0314f8214 New example by Ole Rasmussen
lcp
parents:
diff changeset
     9
datatype
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    10
  "redexes" = Var ("n \<in> nat")            
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    11
            | Fun ("t \<in> redexes")
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 12610
diff changeset
    12
            | App ("b \<in> bool","f \<in> redexes", "a \<in> redexes")
1048
5ba0314f8214 New example by Ole Rasmussen
lcp
parents:
diff changeset
    13
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    14
1048
5ba0314f8214 New example by Ole Rasmussen
lcp
parents:
diff changeset
    15
consts
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    16
  Ssub  :: "i"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    17
  Scomp :: "i"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    18
  Sreg  :: "i"
22808
a7daa74e2980 eliminated unnamed infixes, tuned syntax;
wenzelm
parents: 16417
diff changeset
    19
a7daa74e2980 eliminated unnamed infixes, tuned syntax;
wenzelm
parents: 16417
diff changeset
    20
abbreviation
a7daa74e2980 eliminated unnamed infixes, tuned syntax;
wenzelm
parents: 16417
diff changeset
    21
  Ssub_rel  (infixl "<==" 70) where
a7daa74e2980 eliminated unnamed infixes, tuned syntax;
wenzelm
parents: 16417
diff changeset
    22
  "a<==b == <a,b> \<in> Ssub"
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    23
22808
a7daa74e2980 eliminated unnamed infixes, tuned syntax;
wenzelm
parents: 16417
diff changeset
    24
abbreviation
a7daa74e2980 eliminated unnamed infixes, tuned syntax;
wenzelm
parents: 16417
diff changeset
    25
  Scomp_rel  (infixl "~" 70) where
a7daa74e2980 eliminated unnamed infixes, tuned syntax;
wenzelm
parents: 16417
diff changeset
    26
  "a ~ b == <a,b> \<in> Scomp"
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    27
22808
a7daa74e2980 eliminated unnamed infixes, tuned syntax;
wenzelm
parents: 16417
diff changeset
    28
abbreviation
a7daa74e2980 eliminated unnamed infixes, tuned syntax;
wenzelm
parents: 16417
diff changeset
    29
  "regular(a) == a \<in> Sreg"
9284
85a5355faa91 removal of (harmless) circular definitions
paulson
parents: 6112
diff changeset
    30
22808
a7daa74e2980 eliminated unnamed infixes, tuned syntax;
wenzelm
parents: 16417
diff changeset
    31
consts union_aux        :: "i=>i"
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    32
primrec (*explicit lambda is required because both arguments of "un" vary*)
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    33
  "union_aux(Var(n)) =
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    34
     (\<lambda>t \<in> redexes. redexes_case(%j. Var(n), %x. 0, %b x y.0, t))"
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    35
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    36
  "union_aux(Fun(u)) =
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    37
     (\<lambda>t \<in> redexes. redexes_case(%j. 0, %y. Fun(union_aux(u)`y),
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 22808
diff changeset
    38
                                  %b y z. 0, t))"
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    39
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    40
  "union_aux(App(b,f,a)) =
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    41
     (\<lambda>t \<in> redexes.
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    42
        redexes_case(%j. 0, %y. 0,
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 22808
diff changeset
    43
                     %c z u. App(b or c, union_aux(f)`z, union_aux(a)`u), t))"
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    44
22808
a7daa74e2980 eliminated unnamed infixes, tuned syntax;
wenzelm
parents: 16417
diff changeset
    45
definition
a7daa74e2980 eliminated unnamed infixes, tuned syntax;
wenzelm
parents: 16417
diff changeset
    46
  union  (infixl "un" 70) where
a7daa74e2980 eliminated unnamed infixes, tuned syntax;
wenzelm
parents: 16417
diff changeset
    47
  "u un v == union_aux(u)`v"
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    48
22808
a7daa74e2980 eliminated unnamed infixes, tuned syntax;
wenzelm
parents: 16417
diff changeset
    49
notation (xsymbols)
a7daa74e2980 eliminated unnamed infixes, tuned syntax;
wenzelm
parents: 16417
diff changeset
    50
  union  (infixl "\<squnion>" 70) and
a7daa74e2980 eliminated unnamed infixes, tuned syntax;
wenzelm
parents: 16417
diff changeset
    51
  Ssub_rel  (infixl "\<Longleftarrow>" 70) and
a7daa74e2980 eliminated unnamed infixes, tuned syntax;
wenzelm
parents: 16417
diff changeset
    52
  Scomp_rel  (infixl "\<sim>" 70)
a7daa74e2980 eliminated unnamed infixes, tuned syntax;
wenzelm
parents: 16417
diff changeset
    53
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    54
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    55
inductive
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
    56
  domains       "Ssub" \<subseteq> "redexes*redexes"
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    57
  intros
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    58
    Sub_Var:     "n \<in> nat ==> Var(n)<== Var(n)"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    59
    Sub_Fun:     "[|u<== v|]==> Fun(u)<== Fun(v)"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    60
    Sub_App1:    "[|u1<== v1; u2<== v2; b \<in> bool|]==>   
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    61
                     App(0,u1,u2)<== App(b,v1,v2)"
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    62
    Sub_App2:    "[|u1<== v1; u2<== v2|]==> App(1,u1,u2)<== App(1,v1,v2)"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    63
  type_intros    redexes.intros bool_typechecks
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    64
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    65
inductive
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
    66
  domains       "Scomp" \<subseteq> "redexes*redexes"
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    67
  intros
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    68
    Comp_Var:    "n \<in> nat ==> Var(n) ~ Var(n)"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    69
    Comp_Fun:    "[|u ~ v|]==> Fun(u) ~ Fun(v)"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    70
    Comp_App:    "[|u1 ~ v1; u2 ~ v2; b1 \<in> bool; b2 \<in> bool|]
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    71
                  ==> App(b1,u1,u2) ~ App(b2,v1,v2)"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    72
  type_intros    redexes.intros bool_typechecks
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    73
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    74
inductive
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
    75
  domains       "Sreg" \<subseteq> redexes
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    76
  intros
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    77
    Reg_Var:     "n \<in> nat ==> regular(Var(n))"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    78
    Reg_Fun:     "[|regular(u)|]==> regular(Fun(u))"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    79
    Reg_App1:    "[|regular(Fun(u)); regular(v) |]==>regular(App(1,Fun(u),v))"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    80
    Reg_App2:    "[|regular(u); regular(v) |]==>regular(App(0,u,v))"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    81
  type_intros    redexes.intros bool_typechecks
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    82
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    83
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    84
declare redexes.intros [simp]
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    85
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    86
(* ------------------------------------------------------------------------- *)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    87
(*    Specialisation of comp-rules                                           *)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    88
(* ------------------------------------------------------------------------- *)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    89
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    90
lemmas compD1 [simp] = Scomp.dom_subset [THEN subsetD, THEN SigmaD1]
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    91
lemmas compD2 [simp] = Scomp.dom_subset [THEN subsetD, THEN SigmaD2]
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    92
lemmas regD [simp] = Sreg.dom_subset [THEN subsetD]
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    93
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    94
(* ------------------------------------------------------------------------- *)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    95
(*    Equality rules for union                                               *)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    96
(* ------------------------------------------------------------------------- *)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    97
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    98
lemma union_Var [simp]: "n \<in> nat ==> Var(n) un Var(n)=Var(n)"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    99
by (simp add: union_def)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   100
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   101
lemma union_Fun [simp]: "v \<in> redexes ==> Fun(u) un Fun(v) = Fun(u un v)"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   102
by (simp add: union_def)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   103
 
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   104
lemma union_App [simp]:
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   105
     "[|b2 \<in> bool; u2 \<in> redexes; v2 \<in> redexes|]
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   106
      ==> App(b1,u1,v1) un App(b2,u2,v2)=App(b1 or b2,u1 un u2,v1 un v2)"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   107
by (simp add: union_def)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   108
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   109
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   110
lemma or_1_right [simp]: "a or 1 = 1"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   111
by (simp add: or_def cond_def) 
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   112
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   113
lemma or_0_right [simp]: "a \<in> bool \<Longrightarrow> a or 0 = a"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   114
by (simp add: or_def cond_def bool_def, auto) 
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   115
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   116
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   117
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   118
declare Ssub.intros [simp]
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   119
declare bool_typechecks [simp]
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   120
declare Sreg.intros [simp]
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   121
declare Scomp.intros [simp]
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   122
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   123
declare Scomp.intros [intro]
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
   124
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12593
diff changeset
   125
inductive_cases [elim!]:
8b9845807f77 tuned document sources;
wenzelm
parents: 12593
diff changeset
   126
  "regular(App(b,f,a))"
8b9845807f77 tuned document sources;
wenzelm
parents: 12593
diff changeset
   127
  "regular(Fun(b))"
8b9845807f77 tuned document sources;
wenzelm
parents: 12593
diff changeset
   128
  "regular(Var(b))"
8b9845807f77 tuned document sources;
wenzelm
parents: 12593
diff changeset
   129
  "Fun(u) ~ Fun(t)"
8b9845807f77 tuned document sources;
wenzelm
parents: 12593
diff changeset
   130
  "u ~ Fun(t)"
8b9845807f77 tuned document sources;
wenzelm
parents: 12593
diff changeset
   131
  "u ~ Var(n)"
8b9845807f77 tuned document sources;
wenzelm
parents: 12593
diff changeset
   132
  "u ~ App(b,t,a)"
8b9845807f77 tuned document sources;
wenzelm
parents: 12593
diff changeset
   133
  "Fun(t) ~ v"
8b9845807f77 tuned document sources;
wenzelm
parents: 12593
diff changeset
   134
  "App(b,f,a) ~ v"
8b9845807f77 tuned document sources;
wenzelm
parents: 12593
diff changeset
   135
  "Var(n) ~ u"
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   136
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   137
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   138
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   139
(* ------------------------------------------------------------------------- *)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   140
(*    comp proofs                                                            *)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   141
(* ------------------------------------------------------------------------- *)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   142
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   143
lemma comp_refl [simp]: "u \<in> redexes ==> u ~ u"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   144
by (erule redexes.induct, blast+)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   145
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   146
lemma comp_sym: "u ~ v ==> v ~ u"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   147
by (erule Scomp.induct, blast+)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   148
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   149
lemma comp_sym_iff: "u ~ v \<longleftrightarrow> v ~ u"
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   150
by (blast intro: comp_sym)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   151
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   152
lemma comp_trans [rule_format]: "u ~ v ==> \<forall>w. v ~ w\<longrightarrow>u ~ w"
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   153
by (erule Scomp.induct, blast+)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   154
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   155
(* ------------------------------------------------------------------------- *)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   156
(*   union proofs                                                            *)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   157
(* ------------------------------------------------------------------------- *)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   158
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   159
lemma union_l: "u ~ v ==> u <== (u un v)"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   160
apply (erule Scomp.induct)
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 12610
diff changeset
   161
apply (erule_tac [3] boolE, simp_all)
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   162
done
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   163
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   164
lemma union_r: "u ~ v ==> v <== (u un v)"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   165
apply (erule Scomp.induct)
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 12610
diff changeset
   166
apply (erule_tac [3] c = b2 in boolE, simp_all)
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   167
done
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   168
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   169
lemma union_sym: "u ~ v ==> u un v = v un u"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   170
by (erule Scomp.induct, simp_all add: or_commute)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   171
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   172
(* ------------------------------------------------------------------------- *)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   173
(*      regular proofs                                                       *)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   174
(* ------------------------------------------------------------------------- *)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   175
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   176
lemma union_preserve_regular [rule_format]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   177
     "u ~ v ==> regular(u)\<longrightarrow>regular(v)\<longrightarrow>regular(u un v)"
13612
55d32e76ef4e Adapted to new simplifier.
berghofe
parents: 13339
diff changeset
   178
  by (erule Scomp.induct, auto)
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
   179
1048
5ba0314f8214 New example by Ole Rasmussen
lcp
parents:
diff changeset
   180
end
5ba0314f8214 New example by Ole Rasmussen
lcp
parents:
diff changeset
   181