src/ZF/Resid/Residuals.thy
author wenzelm
Thu, 22 Nov 2007 14:51:34 +0100
changeset 25456 6f79698f294d
parent 24893 b8ef7afe3a6b
child 35762 af3ff2ba4c54
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     1
(*  Title:      Residuals.thy
1048
5ba0314f8214 New example by Ole Rasmussen
lcp
parents:
diff changeset
     2
    ID:         $Id$
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     3
    Author:     Ole Rasmussen
1048
5ba0314f8214 New example by Ole Rasmussen
lcp
parents:
diff changeset
     4
    Copyright   1995  University of Cambridge
5ba0314f8214 New example by Ole Rasmussen
lcp
parents:
diff changeset
     5
    Logic Image: ZF
5ba0314f8214 New example by Ole Rasmussen
lcp
parents:
diff changeset
     6
5ba0314f8214 New example by Ole Rasmussen
lcp
parents:
diff changeset
     7
*)
5ba0314f8214 New example by Ole Rasmussen
lcp
parents:
diff changeset
     8
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 15481
diff changeset
     9
theory Residuals imports Substitution begin
1048
5ba0314f8214 New example by Ole Rasmussen
lcp
parents:
diff changeset
    10
5ba0314f8214 New example by Ole Rasmussen
lcp
parents:
diff changeset
    11
consts
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    12
  Sres          :: "i"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    13
22808
a7daa74e2980 eliminated unnamed infixes, tuned syntax;
wenzelm
parents: 16417
diff changeset
    14
abbreviation
a7daa74e2980 eliminated unnamed infixes, tuned syntax;
wenzelm
parents: 16417
diff changeset
    15
  "residuals(u,v,w) == <u,v,w> \<in> Sres"
1048
5ba0314f8214 New example by Ole Rasmussen
lcp
parents:
diff changeset
    16
5ba0314f8214 New example by Ole Rasmussen
lcp
parents:
diff changeset
    17
inductive
5ba0314f8214 New example by Ole Rasmussen
lcp
parents:
diff changeset
    18
  domains       "Sres" <= "redexes*redexes*redexes"
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    19
  intros
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    20
    Res_Var:    "n \<in> nat ==> residuals(Var(n),Var(n),Var(n))"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    21
    Res_Fun:    "[|residuals(u,v,w)|]==>   
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    22
                     residuals(Fun(u),Fun(v),Fun(w))"
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    23
    Res_App:    "[|residuals(u1,v1,w1);   
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    24
                   residuals(u2,v2,w2); b \<in> bool|]==>   
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    25
                 residuals(App(b,u1,u2),App(0,v1,v2),App(b,w1,w2))"
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    26
    Res_redex:  "[|residuals(u1,v1,w1);   
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    27
                   residuals(u2,v2,w2); b \<in> bool|]==>   
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    28
                 residuals(App(b,Fun(u1),u2),App(1,Fun(v1),v2),w2/w1)"
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    29
  type_intros    subst_type nat_typechecks redexes.intros bool_typechecks
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    30
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 22808
diff changeset
    31
definition
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 22808
diff changeset
    32
  res_func      :: "[i,i]=>i"     (infixl "|>" 70)  where
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 22808
diff changeset
    33
  "u |> v == THE w. residuals(u,v,w)"
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    34
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    35
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 13612
diff changeset
    36
subsection{*Setting up rule lists*}
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    37
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    38
declare Sres.intros [intro]
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    39
declare Sreg.intros [intro]
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    40
declare subst_type [intro]
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    41
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12593
diff changeset
    42
inductive_cases [elim!]:
8b9845807f77 tuned document sources;
wenzelm
parents: 12593
diff changeset
    43
  "residuals(Var(n),Var(n),v)"
8b9845807f77 tuned document sources;
wenzelm
parents: 12593
diff changeset
    44
  "residuals(Fun(t),Fun(u),v)"
8b9845807f77 tuned document sources;
wenzelm
parents: 12593
diff changeset
    45
  "residuals(App(b, u1, u2), App(0, v1, v2),v)"
8b9845807f77 tuned document sources;
wenzelm
parents: 12593
diff changeset
    46
  "residuals(App(b, u1, u2), App(1, Fun(v1), v2),v)"
8b9845807f77 tuned document sources;
wenzelm
parents: 12593
diff changeset
    47
  "residuals(Var(n),u,v)"
8b9845807f77 tuned document sources;
wenzelm
parents: 12593
diff changeset
    48
  "residuals(Fun(t),u,v)"
8b9845807f77 tuned document sources;
wenzelm
parents: 12593
diff changeset
    49
  "residuals(App(b, u1, u2), w,v)"
8b9845807f77 tuned document sources;
wenzelm
parents: 12593
diff changeset
    50
  "residuals(u,Var(n),v)"
8b9845807f77 tuned document sources;
wenzelm
parents: 12593
diff changeset
    51
  "residuals(u,Fun(t),v)"
8b9845807f77 tuned document sources;
wenzelm
parents: 12593
diff changeset
    52
  "residuals(w,App(b, u1, u2),v)"
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    53
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    54
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12593
diff changeset
    55
inductive_cases [elim!]:
8b9845807f77 tuned document sources;
wenzelm
parents: 12593
diff changeset
    56
  "Var(n) <== u"
8b9845807f77 tuned document sources;
wenzelm
parents: 12593
diff changeset
    57
  "Fun(n) <== u"
8b9845807f77 tuned document sources;
wenzelm
parents: 12593
diff changeset
    58
  "u <== Fun(n)"
8b9845807f77 tuned document sources;
wenzelm
parents: 12593
diff changeset
    59
  "App(1,Fun(t),a) <== u"
8b9845807f77 tuned document sources;
wenzelm
parents: 12593
diff changeset
    60
  "App(0,t,a) <== u"
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    61
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12593
diff changeset
    62
inductive_cases [elim!]:
8b9845807f77 tuned document sources;
wenzelm
parents: 12593
diff changeset
    63
  "Fun(t) \<in> redexes"
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    64
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    65
declare Sres.intros [simp]
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    66
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 13612
diff changeset
    67
subsection{*residuals is a  partial function*}
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    68
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    69
lemma residuals_function [rule_format]:
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    70
     "residuals(u,v,w) ==> \<forall>w1. residuals(u,v,w1) --> w1 = w"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    71
by (erule Sres.induct, force+)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    72
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    73
lemma residuals_intro [rule_format]:
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    74
     "u~v ==> regular(v) --> (\<exists>w. residuals(u,v,w))"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    75
by (erule Scomp.induct, force+)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    76
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    77
lemma comp_resfuncD:
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    78
     "[| u~v;  regular(v) |] ==> residuals(u, v, THE w. residuals(u, v, w))"
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 12610
diff changeset
    79
apply (frule residuals_intro, assumption, clarify)
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    80
apply (subst the_equality)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    81
apply (blast intro: residuals_function)+
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    82
done
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    83
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 13612
diff changeset
    84
subsection{*Residual function*}
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    85
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    86
lemma res_Var [simp]: "n \<in> nat ==> Var(n) |> Var(n) = Var(n)"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    87
by (unfold res_func_def, blast)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    88
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    89
lemma res_Fun [simp]: 
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    90
    "[|s~t; regular(t)|]==> Fun(s) |> Fun(t) = Fun(s |> t)"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    91
apply (unfold res_func_def)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    92
apply (blast intro: comp_resfuncD residuals_function) 
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    93
done
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    94
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    95
lemma res_App [simp]: 
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    96
    "[|s~u; regular(u); t~v; regular(v); b \<in> bool|]
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    97
     ==> App(b,s,t) |> App(0,u,v) = App(b, s |> u, t |> v)"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    98
apply (unfold res_func_def) 
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
    99
apply (blast dest!: comp_resfuncD intro: residuals_function)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   100
done
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   101
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   102
lemma res_redex [simp]: 
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   103
    "[|s~u; regular(u); t~v; regular(v); b \<in> bool|]
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   104
     ==> App(b,Fun(s),t) |> App(1,Fun(u),v) = (t |> v)/ (s |> u)"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   105
apply (unfold res_func_def)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   106
apply (blast elim!: redexes.free_elims dest!: comp_resfuncD 
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   107
             intro: residuals_function)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   108
done
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 resfunc_type [simp]:
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   111
     "[|s~t; regular(t)|]==> regular(t) --> s |> t \<in> redexes"
13612
55d32e76ef4e Adapted to new simplifier.
berghofe
parents: 13339
diff changeset
   112
  by (erule Scomp.induct, auto)
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   113
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 13612
diff changeset
   114
subsection{*Commutation theorem*}
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   115
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   116
lemma sub_comp [simp]: "u<==v ==> u~v"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   117
by (erule Ssub.induct, simp_all)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   118
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   119
lemma sub_preserve_reg [rule_format, simp]:
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   120
     "u<==v  ==> regular(v) --> regular(u)"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   121
by (erule Ssub.induct, auto)
1048
5ba0314f8214 New example by Ole Rasmussen
lcp
parents:
diff changeset
   122
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   123
lemma residuals_lift_rec: "[|u~v; k \<in> nat|]==> regular(v)--> (\<forall>n \<in> nat.   
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   124
         lift_rec(u,n) |> lift_rec(v,n) = lift_rec(u |> v,n))"
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 12610
diff changeset
   125
apply (erule Scomp.induct, safe)
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   126
apply (simp_all add: lift_rec_Var subst_Var lift_subst)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   127
done
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   128
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   129
lemma residuals_subst_rec:
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   130
     "u1~u2 ==>  \<forall>v1 v2. v1~v2 --> regular(v2) --> regular(u2) --> 
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   131
                  (\<forall>n \<in> nat. subst_rec(v1,u1,n) |> subst_rec(v2,u2,n) =  
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   132
                    subst_rec(v1 |> v2, u1 |> u2,n))"
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 12610
diff changeset
   133
apply (erule Scomp.induct, safe)
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   134
apply (simp_all add: lift_rec_Var subst_Var residuals_lift_rec)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   135
apply (drule_tac psi = "\<forall>x.?P (x) " in asm_rl)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   136
apply (simp add: substitution)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   137
done
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
lemma commutation [simp]:
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   141
     "[|u1~u2; v1~v2; regular(u2); regular(v2)|]
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   142
      ==> (v1/u1) |> (v2/u2) = (v1 |> v2)/(u1 |> u2)"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   143
by (simp add: residuals_subst_rec)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   144
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   145
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 13612
diff changeset
   146
subsection{*Residuals are comp and regular*}
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   147
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   148
lemma residuals_preserve_comp [rule_format, simp]:
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   149
     "u~v ==> \<forall>w. u~w --> v~w --> regular(w) --> (u|>w) ~ (v|>w)"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   150
by (erule Scomp.induct, force+)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   151
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   152
lemma residuals_preserve_reg [rule_format, simp]:
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   153
     "u~v ==> regular(u) --> regular(v) --> regular(u|>v)"
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 12610
diff changeset
   154
apply (erule Scomp.induct, auto)
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   155
done
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   156
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 13612
diff changeset
   157
subsection{*Preservation lemma*}
12593
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_preserve_comp: "u~v ==> v ~ (u un v)"
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   160
by (erule Scomp.induct, simp_all)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   161
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   162
lemma preservation [rule_format]:
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   163
     "u ~ v ==> regular(v) --> u|>v = (u un v)|>v"
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 12610
diff changeset
   164
apply (erule Scomp.induct, safe)
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   165
apply (drule_tac [3] psi = "Fun (?u) |> ?v = ?w" in asm_rl)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   166
apply (auto simp add: union_preserve_comp comp_sym_iff)
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
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   170
declare sub_comp [THEN comp_sym, simp]
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   171
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 13612
diff changeset
   172
subsection{*Prism theorem*}
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   173
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   174
(* Having more assumptions than needed -- removed below  *)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   175
lemma prism_l [rule_format]:
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   176
     "v<==u ==>  
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   177
       regular(u) --> (\<forall>w. w~v --> w~u -->   
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   178
                            w |> u = (w|>v) |> (u|>v))"
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 13612
diff changeset
   179
by (erule Ssub.induct, force+)
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   180
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 13612
diff changeset
   181
lemma prism: "[|v <== u; regular(u); w~v|] ==> w |> u = (w|>v) |> (u|>v)"
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   182
apply (rule prism_l)
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 12610
diff changeset
   183
apply (rule_tac [4] comp_trans, auto)
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   184
done
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   185
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   186
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 13612
diff changeset
   187
subsection{*Levy's Cube Lemma*}
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   188
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   189
lemma cube: "[|u~v; regular(v); regular(u); w~u|]==>   
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   190
           (w|>u) |> (v|>u) = (w|>v) |> (u|>v)"
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 13612
diff changeset
   191
apply (subst preservation [of u], assumption, assumption)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 13612
diff changeset
   192
apply (subst preservation [of v], erule comp_sym, assumption)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 13612
diff changeset
   193
apply (subst prism [symmetric, of v])
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   194
apply (simp add: union_r comp_sym_iff)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   195
apply (simp add: union_preserve_regular comp_sym_iff)
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 12610
diff changeset
   196
apply (erule comp_trans, assumption)
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   197
apply (simp add: prism [symmetric] union_l union_preserve_regular 
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   198
                 comp_sym_iff union_sym)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   199
done
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   200
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   201
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 13612
diff changeset
   202
subsection{*paving theorem*}
12593
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   203
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   204
lemma paving: "[|w~u; w~v; regular(u); regular(v)|]==>  
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   205
           \<exists>uv vu. (w|>u) |> vu = (w|>v) |> uv & (w|>u)~vu & 
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   206
             regular(vu) & (w|>v)~uv & regular(uv) "
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   207
apply (subgoal_tac "u~v")
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   208
apply (safe intro!: exI)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   209
apply (rule cube)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   210
apply (simp_all add: comp_sym_iff)
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   211
apply (blast intro: residuals_preserve_comp comp_trans comp_sym)+
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   212
done
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   213
cd35fe5947d4 Resid converted to Isar/ZF
paulson
parents: 11319
diff changeset
   214
1048
5ba0314f8214 New example by Ole Rasmussen
lcp
parents:
diff changeset
   215
end
5ba0314f8214 New example by Ole Rasmussen
lcp
parents:
diff changeset
   216