src/HOL/Induct/Comb.thy
author wenzelm
Wed, 28 Dec 2022 12:30:18 +0100
changeset 76798 69d8d16c5612
parent 71591 8e4d542f041b
child 76987 4c275405faae
permissions -rw-r--r--
tuned output;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
     1
(*  Title:      HOL/Induct/Comb.thy
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     2
    Author:     Lawrence C Paulson
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     3
    Copyright   1996  University of Cambridge
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     4
*)
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     5
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
     6
section \<open>Combinatory Logic example: the Church-Rosser Theorem\<close>
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     7
62045
75a7db3cae7e more symbols;
wenzelm
parents: 61993
diff changeset
     8
theory Comb
75a7db3cae7e more symbols;
wenzelm
parents: 61993
diff changeset
     9
imports Main
75a7db3cae7e more symbols;
wenzelm
parents: 61993
diff changeset
    10
begin
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
    11
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    12
text \<open>
71591
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    13
  Combinator terms do not have free variables.
58623
2db1df2c8467 more bibtex entries;
wenzelm
parents: 58310
diff changeset
    14
  Example taken from @{cite camilleri92}.
62045
75a7db3cae7e more symbols;
wenzelm
parents: 61993
diff changeset
    15
\<close>
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    16
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    17
subsection \<open>Definitions\<close>
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
    18
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    19
text \<open>Datatype definition of combinators \<open>S\<close> and \<open>K\<close>.\<close>
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
    20
58310
91ea607a34d8 updated news
blanchet
parents: 58249
diff changeset
    21
datatype comb = K
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    22
              | S
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    23
              | Ap comb comb (infixl "\<bullet>" 90)
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    24
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    25
text \<open>
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    26
  Inductive definition of contractions, \<open>\<rightarrow>\<^sup>1\<close> and
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    27
  (multi-step) reductions, \<open>\<rightarrow>\<close>.
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    28
\<close>
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
    29
71591
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    30
inductive contract1 :: "[comb,comb] \<Rightarrow> bool"  (infixl "\<rightarrow>\<^sup>1" 50)
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    31
  where
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    32
    K:     "K\<bullet>x\<bullet>y \<rightarrow>\<^sup>1 x"
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    33
  | S:     "S\<bullet>x\<bullet>y\<bullet>z \<rightarrow>\<^sup>1 (x\<bullet>z)\<bullet>(y\<bullet>z)"
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    34
  | Ap1:   "x \<rightarrow>\<^sup>1 y \<Longrightarrow> x\<bullet>z \<rightarrow>\<^sup>1 y\<bullet>z"
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    35
  | Ap2:   "x \<rightarrow>\<^sup>1 y \<Longrightarrow> z\<bullet>x \<rightarrow>\<^sup>1 z\<bullet>y"
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21210
diff changeset
    36
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21210
diff changeset
    37
abbreviation
71591
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    38
  contract :: "[comb,comb] \<Rightarrow> bool"   (infixl "\<rightarrow>" 50) where
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    39
  "contract \<equiv> contract1\<^sup>*\<^sup>*"
15816
4575c87dd25b x-symbol syntax
paulson
parents: 13075
diff changeset
    40
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    41
text \<open>
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    42
  Inductive definition of parallel contractions, \<open>\<Rrightarrow>\<^sup>1\<close> and
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    43
  (multi-step) parallel reductions, \<open>\<Rrightarrow>\<close>.
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    44
\<close>
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    45
71591
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    46
inductive parcontract1 :: "[comb,comb] \<Rightarrow> bool"  (infixl "\<Rrightarrow>\<^sup>1" 50)
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    47
  where
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    48
    refl:  "x \<Rrightarrow>\<^sup>1 x"
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    49
  | K:     "K\<bullet>x\<bullet>y \<Rrightarrow>\<^sup>1 x"
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    50
  | S:     "S\<bullet>x\<bullet>y\<bullet>z \<Rrightarrow>\<^sup>1 (x\<bullet>z)\<bullet>(y\<bullet>z)"
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    51
  | Ap:    "\<lbrakk>x \<Rrightarrow>\<^sup>1 y; z \<Rrightarrow>\<^sup>1 w\<rbrakk> \<Longrightarrow> x\<bullet>z \<Rrightarrow>\<^sup>1 y\<bullet>w"
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21210
diff changeset
    52
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21210
diff changeset
    53
abbreviation
71591
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    54
  parcontract :: "[comb,comb] \<Rightarrow> bool"   (infixl "\<Rrightarrow>" 50) where
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    55
  "parcontract \<equiv> parcontract1\<^sup>*\<^sup>*"
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    56
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    57
text \<open>
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
    58
  Misc definitions.
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    59
\<close>
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    60
19736
wenzelm
parents: 16588
diff changeset
    61
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21210
diff changeset
    62
  I :: comb where
71591
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    63
  "I \<equiv> S\<bullet>K\<bullet>K"
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    64
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21210
diff changeset
    65
definition
71591
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    66
  diamond   :: "([comb,comb] \<Rightarrow> bool) \<Rightarrow> bool" where
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 62045
diff changeset
    67
    \<comment> \<open>confluence; Lambda/Commutation treats this more abstractly\<close>
71591
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    68
  "diamond r \<equiv> \<forall>x y. r x y \<longrightarrow>
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    69
                  (\<forall>y'. r x y' \<longrightarrow> 
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    70
                    (\<exists>z. r y z \<and> r y' z))"
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
    71
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
    72
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    73
subsection \<open>Reflexive/Transitive closure preserves Church-Rosser property\<close>
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
    74
71591
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    75
text\<open>Remark: So does the Transitive closure, with a similar proof\<close>
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
    76
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    77
text\<open>Strip lemma.  
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    78
   The induction hypothesis covers all but the last diamond of the strip.\<close>
71591
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    79
lemma strip_lemma [rule_format]: 
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    80
  assumes "diamond r" and r: "r\<^sup>*\<^sup>* x y" "r x y'"
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    81
  shows "\<exists>z. r\<^sup>*\<^sup>* y' z \<and> r y z"
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    82
  using r
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    83
proof (induction rule: rtranclp_induct)
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    84
  case base
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    85
  then show ?case
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    86
    by blast
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    87
next
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    88
  case (step y z)
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    89
  then show ?case
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    90
    using \<open>diamond r\<close> unfolding diamond_def
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    91
    by (metis rtranclp.rtrancl_into_rtrancl)
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    92
qed
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
    93
71591
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    94
proposition diamond_rtrancl:
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    95
  assumes "diamond r" 
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    96
  shows "diamond(r\<^sup>*\<^sup>*)"
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    97
  unfolding diamond_def
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    98
proof (intro strip)
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    99
  fix x y y'
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   100
  assume "r\<^sup>*\<^sup>* x y" "r\<^sup>*\<^sup>* x y'"
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   101
  then show "\<exists>z. r\<^sup>*\<^sup>* y z \<and> r\<^sup>*\<^sup>* y' z"
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   102
  proof (induction rule: rtranclp_induct)
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   103
    case base
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   104
    then show ?case
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   105
      by blast
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   106
  next
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   107
    case (step y z)
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   108
    then show ?case
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   109
      by (meson assms strip_lemma rtranclp.rtrancl_into_rtrancl)
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   110
  qed
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   111
qed
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   112
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   113
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   114
subsection \<open>Non-contraction results\<close>
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   115
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   116
text \<open>Derive a case for each combinator constructor.\<close>
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   117
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   118
inductive_cases
71591
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   119
  K_contractE [elim!]: "K \<rightarrow>\<^sup>1 r"
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
   120
  and S_contractE [elim!]: "S \<rightarrow>\<^sup>1 r"
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
   121
  and Ap_contractE [elim!]: "p\<bullet>q \<rightarrow>\<^sup>1 r"
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   122
71591
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   123
declare contract1.K [intro!] contract1.S [intro!]
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   124
declare contract1.Ap1 [intro] contract1.Ap2 [intro]
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   125
71591
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   126
lemma I_contract_E [iff]: "\<not> I \<rightarrow>\<^sup>1 z"
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   127
  unfolding I_def by blast
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   128
71591
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   129
lemma K1_contractD [elim!]: "K\<bullet>x \<rightarrow>\<^sup>1 z \<Longrightarrow> (\<exists>x'. z = K\<bullet>x' \<and> x \<rightarrow>\<^sup>1 x')"
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   130
  by blast
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   131
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
   132
lemma Ap_reduce1 [intro]: "x \<rightarrow> y \<Longrightarrow> x\<bullet>z \<rightarrow> y\<bullet>z"
71591
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   133
  by (induction rule: rtranclp_induct; blast intro: rtranclp_trans)
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   134
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
   135
lemma Ap_reduce2 [intro]: "x \<rightarrow> y \<Longrightarrow> z\<bullet>x \<rightarrow> z\<bullet>y"
71591
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   136
  by (induction rule: rtranclp_induct; blast intro: rtranclp_trans)
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   137
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
   138
text \<open>Counterexample to the diamond property for \<^term>\<open>x \<rightarrow>\<^sup>1 y\<close>\<close>
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   139
71591
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   140
lemma not_diamond_contract: "\<not> diamond(contract1)"
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   141
  unfolding diamond_def by (metis S_contractE contract1.K) 
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   142
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   143
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   144
subsection \<open>Results about Parallel Contraction\<close>
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   145
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   146
text \<open>Derive a case for each combinator constructor.\<close>
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   147
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   148
inductive_cases
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
   149
      K_parcontractE [elim!]: "K \<Rrightarrow>\<^sup>1 r"
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
   150
  and S_parcontractE [elim!]: "S \<Rrightarrow>\<^sup>1 r"
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
   151
  and Ap_parcontractE [elim!]: "p\<bullet>q \<Rrightarrow>\<^sup>1 r"
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   152
71591
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   153
declare parcontract1.intros [intro]
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   154
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   155
subsection \<open>Basic properties of parallel contraction\<close>
71591
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   156
text\<open>The rules below are not essential but make proofs much faster\<close>
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   157
71591
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   158
lemma K1_parcontractD [dest!]: "K\<bullet>x \<Rrightarrow>\<^sup>1 z \<Longrightarrow> (\<exists>x'. z = K\<bullet>x' \<and> x \<Rrightarrow>\<^sup>1 x')"
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   159
  by blast
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   160
71591
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   161
lemma S1_parcontractD [dest!]: "S\<bullet>x \<Rrightarrow>\<^sup>1 z \<Longrightarrow> (\<exists>x'. z = S\<bullet>x' \<and> x \<Rrightarrow>\<^sup>1 x')"
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   162
  by blast
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   163
71591
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   164
lemma S2_parcontractD [dest!]: "S\<bullet>x\<bullet>y \<Rrightarrow>\<^sup>1 z \<Longrightarrow> (\<exists>x' y'. z = S\<bullet>x'\<bullet>y' \<and> x \<Rrightarrow>\<^sup>1 x' \<and> y \<Rrightarrow>\<^sup>1 y')"
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   165
  by blast
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   166
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   167
text\<open>Church-Rosser property for parallel contraction\<close>
71591
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   168
proposition diamond_parcontract: "diamond parcontract1"
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   169
proof -
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   170
  have "(\<exists>z. w \<Rrightarrow>\<^sup>1 z \<and> y' \<Rrightarrow>\<^sup>1 z)" if "y \<Rrightarrow>\<^sup>1 w" "y \<Rrightarrow>\<^sup>1 y'" for w y y'
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   171
    using that by (induction arbitrary: y' rule: parcontract1.induct) fast+
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   172
  then show ?thesis
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   173
    by (auto simp: diamond_def)
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   174
qed
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   175
71591
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   176
subsection \<open>Equivalence of \<^prop>\<open>p \<rightarrow> q\<close> and \<^prop>\<open>p \<Rrightarrow> q\<close>.\<close>
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   177
71591
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   178
lemma contract_imp_parcontract: "x \<rightarrow>\<^sup>1 y \<Longrightarrow> x \<Rrightarrow>\<^sup>1 y"
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   179
  by (induction rule: contract1.induct; blast)
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   180
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   181
text\<open>Reductions: simply throw together reflexivity, transitivity and
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   182
  the one-step reductions\<close>
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   183
71591
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   184
proposition reduce_I: "I\<bullet>x \<rightarrow> x"
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   185
  unfolding I_def
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   186
  by (meson contract1.K contract1.S r_into_rtranclp rtranclp.rtrancl_into_rtrancl)
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   187
71591
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   188
lemma parcontract_imp_reduce: "x \<Rrightarrow>\<^sup>1 y \<Longrightarrow> x \<rightarrow> y"
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   189
proof (induction rule: parcontract1.induct)
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   190
  case (Ap x y z w)
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   191
  then show ?case
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   192
    by (meson Ap_reduce1 Ap_reduce2 rtranclp_trans)
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   193
qed auto
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   194
71591
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   195
lemma reduce_eq_parreduce: "x \<rightarrow> y  \<longleftrightarrow>  x \<Rrightarrow> y"
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   196
  by (metis contract_imp_parcontract parcontract_imp_reduce predicate2I rtranclp_subset)
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   197
71591
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   198
theorem diamond_reduce: "diamond(contract)"
8e4d542f041b updated to more modern style
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   199
  using diamond_parcontract diamond_rtrancl reduce_eq_parreduce by presburger
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   200
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   201
end