src/ZF/ex/Commutation.thy
author kleing
Wed, 07 Jan 2004 07:52:12 +0100
changeset 14343 6bc647f472b9
parent 13339 0f89104dd377
child 16417 9bc16273c2d4
permissions -rw-r--r--
map_idI
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11042
bb566dd3f927 commutation theory, ported by Sidi Ehmety
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Lambda/Commutation.thy
bb566dd3f927 commutation theory, ported by Sidi Ehmety
paulson
parents:
diff changeset
     2
    ID:         $Id$
bb566dd3f927 commutation theory, ported by Sidi Ehmety
paulson
parents:
diff changeset
     3
    Author:     Tobias Nipkow & Sidi Ould Ehmety
bb566dd3f927 commutation theory, ported by Sidi Ehmety
paulson
parents:
diff changeset
     4
    Copyright   1995  TU Muenchen
bb566dd3f927 commutation theory, ported by Sidi Ehmety
paulson
parents:
diff changeset
     5
bb566dd3f927 commutation theory, ported by Sidi Ehmety
paulson
parents:
diff changeset
     6
Commutation theory for proving the Church Rosser theorem
bb566dd3f927 commutation theory, ported by Sidi Ehmety
paulson
parents:
diff changeset
     7
	ported from Isabelle/HOL  by Sidi Ould Ehmety
bb566dd3f927 commutation theory, ported by Sidi Ehmety
paulson
parents:
diff changeset
     8
*)
bb566dd3f927 commutation theory, ported by Sidi Ehmety
paulson
parents:
diff changeset
     9
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    10
theory Commutation = Main:
11042
bb566dd3f927 commutation theory, ported by Sidi Ehmety
paulson
parents:
diff changeset
    11
bb566dd3f927 commutation theory, ported by Sidi Ehmety
paulson
parents:
diff changeset
    12
constdefs
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    13
  square  :: "[i, i, i, i] => o"
11042
bb566dd3f927 commutation theory, ported by Sidi Ehmety
paulson
parents:
diff changeset
    14
   "square(r,s,t,u) ==
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    15
     (\<forall>a b. <a,b> \<in> r --> (\<forall>c. <a, c> \<in> s --> (\<exists>x. <b,x> \<in> t & <c,x> \<in> u)))"
11042
bb566dd3f927 commutation theory, ported by Sidi Ehmety
paulson
parents:
diff changeset
    16
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    17
  commute :: "[i, i] => o"
11042
bb566dd3f927 commutation theory, ported by Sidi Ehmety
paulson
parents:
diff changeset
    18
   "commute(r,s) == square(r,s,s,r)"
bb566dd3f927 commutation theory, ported by Sidi Ehmety
paulson
parents:
diff changeset
    19
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    20
  diamond :: "i=>o"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    21
   "diamond(r)   == commute(r, r)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    22
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    23
  strip :: "i=>o"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    24
   "strip(r) == commute(r^*, r)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    25
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    26
  Church_Rosser :: "i => o"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    27
   "Church_Rosser(r) == (\<forall>x y. <x,y> \<in>  (r Un converse(r))^* -->
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    28
	 	 	 (\<exists>z. <x,z> \<in> r^* & <y,z> \<in> r^*))"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    29
  confluent :: "i=>o"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    30
   "confluent(r) == diamond(r^*)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    31
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    32
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    33
lemma square_sym: "square(r,s,t,u) ==> square(s,r,u,t)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    34
by (unfold square_def, blast)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    35
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    36
lemma square_subset: "[| square(r,s,t,u); t \<subseteq> t' |] ==> square(r,s,t',u)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    37
by (unfold square_def, blast)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    38
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    39
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    40
lemma square_rtrancl [rule_format]: 
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    41
     "square(r,s,s,t) --> field(s)<=field(t) --> square(r^*,s,s,t^*)"
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13248
diff changeset
    42
apply (unfold square_def, clarify)
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    43
apply (erule rtrancl_induct)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    44
apply (blast intro: rtrancl_refl)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    45
apply (blast intro: rtrancl_into_rtrancl)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    46
done
11042
bb566dd3f927 commutation theory, ported by Sidi Ehmety
paulson
parents:
diff changeset
    47
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    48
(* A special case of square_rtrancl_on *)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    49
lemma diamond_strip: 
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    50
 "diamond(r) ==> strip(r)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    51
apply (unfold diamond_def commute_def strip_def)
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13248
diff changeset
    52
apply (rule square_rtrancl, simp_all)
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    53
done
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    54
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    55
(*** commute ***)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    56
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    57
lemma commute_sym: 
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    58
    "commute(r,s) ==> commute(s,r)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    59
by (unfold commute_def, blast intro: square_sym)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    60
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    61
lemma commute_rtrancl [rule_format]: 
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    62
    "commute(r,s) ==> field(r)=field(s) --> commute(r^*,s^*)"
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13248
diff changeset
    63
apply (unfold commute_def, clarify)
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    64
apply (rule square_rtrancl)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    65
apply (rule square_sym [THEN square_rtrancl, THEN square_sym]) 
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    66
apply (simp_all add: rtrancl_field)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    67
done
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    68
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    69
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    70
lemma confluentD: "confluent(r) ==> diamond(r^*)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    71
by (simp add: confluent_def)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    72
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    73
lemma strip_confluent: "strip(r) ==> confluent(r)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    74
apply (unfold strip_def confluent_def diamond_def)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    75
apply (drule commute_rtrancl)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    76
apply (simp_all add: rtrancl_field)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    77
done
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    78
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    79
lemma commute_Un: "[| commute(r,t); commute(s,t) |] ==> commute(r Un s, t)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    80
by (unfold commute_def square_def, blast)
11042
bb566dd3f927 commutation theory, ported by Sidi Ehmety
paulson
parents:
diff changeset
    81
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    82
lemma diamond_Un: 
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    83
     "[| diamond(r); diamond(s); commute(r, s) |] ==> diamond(r Un s)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    84
by (unfold diamond_def, blast intro: commute_Un commute_sym) 
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    85
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    86
lemma diamond_confluent: 
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    87
    "diamond(r) ==> confluent(r)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    88
apply (unfold diamond_def confluent_def)
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13248
diff changeset
    89
apply (erule commute_rtrancl, simp)
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    90
done
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    91
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    92
lemma confluent_Un: 
13248
ae66c22ed52e new theorems
paulson
parents: 12867
diff changeset
    93
 "[| confluent(r); confluent(s); commute(r^*, s^*); 
ae66c22ed52e new theorems
paulson
parents: 12867
diff changeset
    94
     relation(r); relation(s) |] ==> confluent(r Un s)"
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    95
apply (unfold confluent_def)
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13248
diff changeset
    96
apply (rule rtrancl_Un_rtrancl [THEN subst], auto)
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    97
apply (blast dest: diamond_Un intro: diamond_confluent [THEN confluentD])
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    98
done
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    99
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   100
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   101
lemma diamond_to_confluence: 
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   102
     "[| diamond(r); s \<subseteq> r; r<= s^* |] ==> confluent(s)"
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13248
diff changeset
   103
apply (drule rtrancl_subset [symmetric], assumption)
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   104
apply (simp_all add: confluent_def)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   105
apply (blast intro: diamond_confluent [THEN confluentD])
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   106
done
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   107
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   108
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   109
(*** Church_Rosser ***)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   110
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   111
lemma Church_Rosser1: 
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   112
     "Church_Rosser(r) ==> confluent(r)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   113
apply (unfold confluent_def Church_Rosser_def square_def 
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13248
diff changeset
   114
              commute_def diamond_def, auto)
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   115
apply (drule converseI)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   116
apply (simp (no_asm_use) add: rtrancl_converse [symmetric])
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13248
diff changeset
   117
apply (drule_tac x = b in spec)
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13248
diff changeset
   118
apply (drule_tac x1 = c in spec [THEN mp])
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13248
diff changeset
   119
apply (rule_tac b = a in rtrancl_trans)
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   120
apply (blast intro: rtrancl_mono [THEN subsetD])+
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   121
done
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   122
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   123
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   124
lemma Church_Rosser2: 
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   125
     "confluent(r) ==> Church_Rosser(r)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   126
apply (unfold confluent_def Church_Rosser_def square_def 
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13248
diff changeset
   127
              commute_def diamond_def, auto)
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   128
apply (frule fieldI1)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   129
apply (simp add: rtrancl_field)
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13248
diff changeset
   130
apply (erule rtrancl_induct, auto)
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   131
apply (blast intro: rtrancl_refl)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   132
apply (blast del: rtrancl_refl intro: r_into_rtrancl rtrancl_trans)+
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   133
done
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   134
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   135
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   136
lemma Church_Rosser: "Church_Rosser(r) <-> confluent(r)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   137
by (blast intro: Church_Rosser1 Church_Rosser2)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   138
11042
bb566dd3f927 commutation theory, ported by Sidi Ehmety
paulson
parents:
diff changeset
   139
end