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