src/ZF/ex/Commutation.thy
author wenzelm
Sun, 09 Apr 2017 20:44:35 +0200
changeset 65449 c82e63b11b8b
parent 46822 95f1e700b712
child 76213 e44d86131648
permissions -rw-r--r--
clarified main ZF.thy / ZFC.thy, and avoid name clash with global HOL/Main.thy;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
41777
1f7cbe39d425 more precise headers;
wenzelm
parents: 35762
diff changeset
     1
(*  Title:      ZF/ex/Commutation.thy
11042
bb566dd3f927 commutation theory, ported by Sidi Ehmety
paulson
parents:
diff changeset
     2
    Author:     Tobias Nipkow & Sidi Ould Ehmety
bb566dd3f927 commutation theory, ported by Sidi Ehmety
paulson
parents:
diff changeset
     3
    Copyright   1995  TU Muenchen
bb566dd3f927 commutation theory, ported by Sidi Ehmety
paulson
parents:
diff changeset
     4
35762
af3ff2ba4c54 removed old CVS Ids;
wenzelm
parents: 21404
diff changeset
     5
Commutation theory for proving the Church Rosser theorem.
11042
bb566dd3f927 commutation theory, ported by Sidi Ehmety
paulson
parents:
diff changeset
     6
*)
bb566dd3f927 commutation theory, ported by Sidi Ehmety
paulson
parents:
diff changeset
     7
65449
c82e63b11b8b clarified main ZF.thy / ZFC.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 46822
diff changeset
     8
theory Commutation imports ZF begin
11042
bb566dd3f927 commutation theory, ported by Sidi Ehmety
paulson
parents:
diff changeset
     9
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
    10
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    11
  square  :: "[i, i, i, i] => o" where
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
    12
  "square(r,s,t,u) ==
46822
95f1e700b712 mathematical symbols for Isabelle/ZF example theories
paulson
parents: 41777
diff changeset
    13
    (\<forall>a b. <a,b> \<in> r \<longrightarrow> (\<forall>c. <a, c> \<in> s \<longrightarrow> (\<exists>x. <b,x> \<in> t & <c,x> \<in> u)))"
11042
bb566dd3f927 commutation theory, ported by Sidi Ehmety
paulson
parents:
diff changeset
    14
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    15
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    16
  commute :: "[i, i] => o" where
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
    17
  "commute(r,s) == square(r,s,s,r)"
11042
bb566dd3f927 commutation theory, ported by Sidi Ehmety
paulson
parents:
diff changeset
    18
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    19
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    20
  diamond :: "i=>o" where
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
    21
  "diamond(r)   == commute(r, r)"
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    22
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    23
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    24
  strip :: "i=>o" where
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
    25
  "strip(r) == commute(r^*, r)"
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    26
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    27
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    28
  Church_Rosser :: "i => o" where
46822
95f1e700b712 mathematical symbols for Isabelle/ZF example theories
paulson
parents: 41777
diff changeset
    29
  "Church_Rosser(r) == (\<forall>x y. <x,y> \<in>  (r \<union> converse(r))^* \<longrightarrow>
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
    30
                        (\<exists>z. <x,z> \<in> r^* & <y,z> \<in> r^*))"
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    31
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    32
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    33
  confluent :: "i=>o" where
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
    34
  "confluent(r) == diamond(r^*)"
12867
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
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    37
lemma square_sym: "square(r,s,t,u) ==> square(s,r,u,t)"
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
    38
  unfolding square_def by blast
12867
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_subset: "[| square(r,s,t,u); t \<subseteq> t' |] ==> square(r,s,t',u)"
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
    41
  unfolding square_def by blast
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    42
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    43
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
    44
lemma square_rtrancl:
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
    45
  "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
    46
apply (unfold square_def, clarify)
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    47
apply (erule rtrancl_induct)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    48
apply (blast intro: rtrancl_refl)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    49
apply (blast intro: rtrancl_into_rtrancl)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    50
done
11042
bb566dd3f927 commutation theory, ported by Sidi Ehmety
paulson
parents:
diff changeset
    51
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    52
(* A special case of square_rtrancl_on *)
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
    53
lemma diamond_strip:
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
    54
  "diamond(r) ==> strip(r)"
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    55
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
    56
apply (rule square_rtrancl, simp_all)
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    57
done
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    58
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    59
(*** commute ***)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    60
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
    61
lemma commute_sym: "commute(r,s) ==> commute(s,r)"
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
    62
  unfolding commute_def by (blast intro: square_sym)
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    63
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
    64
lemma commute_rtrancl:
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
    65
  "commute(r,s) ==> field(r)=field(s) ==> commute(r^*,s^*)"
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
    66
apply (unfold commute_def)
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    67
apply (rule square_rtrancl)
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
    68
apply (rule square_sym [THEN square_rtrancl, THEN square_sym])
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    69
apply (simp_all add: rtrancl_field)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    70
done
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    71
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 confluentD: "confluent(r) ==> diamond(r^*)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    74
by (simp add: confluent_def)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    75
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    76
lemma strip_confluent: "strip(r) ==> confluent(r)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    77
apply (unfold strip_def confluent_def diamond_def)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    78
apply (drule commute_rtrancl)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    79
apply (simp_all add: rtrancl_field)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    80
done
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    81
46822
95f1e700b712 mathematical symbols for Isabelle/ZF example theories
paulson
parents: 41777
diff changeset
    82
lemma commute_Un: "[| commute(r,t); commute(s,t) |] ==> commute(r \<union> s, t)"
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
    83
  unfolding commute_def square_def by blast
11042
bb566dd3f927 commutation theory, ported by Sidi Ehmety
paulson
parents:
diff changeset
    84
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
    85
lemma diamond_Un:
46822
95f1e700b712 mathematical symbols for Isabelle/ZF example theories
paulson
parents: 41777
diff changeset
    86
     "[| diamond(r); diamond(s); commute(r, s) |] ==> diamond(r \<union> s)"
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
    87
  unfolding diamond_def by (blast intro: commute_Un commute_sym)
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    88
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
    89
lemma diamond_confluent:
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    90
    "diamond(r) ==> confluent(r)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    91
apply (unfold diamond_def confluent_def)
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13248
diff changeset
    92
apply (erule commute_rtrancl, simp)
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    93
done
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    94
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
    95
lemma confluent_Un:
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
    96
 "[| confluent(r); confluent(s); commute(r^*, s^*);
46822
95f1e700b712 mathematical symbols for Isabelle/ZF example theories
paulson
parents: 41777
diff changeset
    97
     relation(r); relation(s) |] ==> confluent(r \<union> s)"
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
    98
apply (unfold confluent_def)
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13248
diff changeset
    99
apply (rule rtrancl_Un_rtrancl [THEN subst], auto)
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   100
apply (blast dest: diamond_Un intro: diamond_confluent [THEN confluentD])
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   101
done
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   102
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   103
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
   104
lemma diamond_to_confluence:
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   105
     "[| 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
   106
apply (drule rtrancl_subset [symmetric], assumption)
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   107
apply (simp_all add: confluent_def)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   108
apply (blast intro: diamond_confluent [THEN confluentD])
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   109
done
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
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   112
(*** Church_Rosser ***)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   113
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
   114
lemma Church_Rosser1:
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   115
     "Church_Rosser(r) ==> confluent(r)"
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
   116
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
   117
              commute_def diamond_def, auto)
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   118
apply (drule converseI)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   119
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
   120
apply (drule_tac x = b in spec)
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13248
diff changeset
   121
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
   122
apply (rule_tac b = a in rtrancl_trans)
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   123
apply (blast intro: rtrancl_mono [THEN subsetD])+
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   124
done
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   125
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   126
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
   127
lemma Church_Rosser2:
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   128
     "confluent(r) ==> Church_Rosser(r)"
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
   129
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
   130
              commute_def diamond_def, auto)
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   131
apply (frule fieldI1)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   132
apply (simp add: rtrancl_field)
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13248
diff changeset
   133
apply (erule rtrancl_induct, auto)
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   134
apply (blast intro: rtrancl_refl)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   135
apply (blast del: rtrancl_refl intro: r_into_rtrancl rtrancl_trans)+
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   136
done
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   137
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   138
46822
95f1e700b712 mathematical symbols for Isabelle/ZF example theories
paulson
parents: 41777
diff changeset
   139
lemma Church_Rosser: "Church_Rosser(r) \<longleftrightarrow> confluent(r)"
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
   140
  by (blast intro: Church_Rosser1 Church_Rosser2)
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11316
diff changeset
   141
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
   142
end