src/HOL/Induct/Comb.thy
author wenzelm
Wed, 30 Dec 2015 20:24:43 +0100
changeset 61993 89206877f0ee
parent 60530 44f9873d6f6f
child 62045 75a7db3cae7e
permissions -rw-r--r--
more symbols;
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
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 15816
diff changeset
     8
theory Comb imports Main begin
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
     9
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    10
text \<open>
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
    11
  Curiously, combinators do not include free variables.
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
    12
58623
2db1df2c8467 more bibtex entries;
wenzelm
parents: 58310
diff changeset
    13
  Example taken from @{cite camilleri92}.
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    14
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
    15
HOL system proofs may be found in the HOL distribution at
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
    16
   .../contrib/rule-induction/cl.ml
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    17
\<close>
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
    18
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    19
subsection \<open>Definitions\<close>
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
    20
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    21
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
    22
58310
91ea607a34d8 updated news
blanchet
parents: 58249
diff changeset
    23
datatype comb = K
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    24
              | S
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    25
              | Ap comb comb (infixl "\<bullet>" 90)
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    26
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    27
text \<open>
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    28
  Inductive definition of contractions, \<open>\<rightarrow>\<^sup>1\<close> and
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    29
  (multi-step) reductions, \<open>\<rightarrow>\<close>.
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    30
\<close>
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
    31
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    32
inductive_set contract :: "(comb*comb) set"
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    33
  and contract_rel1 :: "[comb,comb] \<Rightarrow> bool"  (infixl "\<rightarrow>\<^sup>1" 50)
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    34
where
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    35
  "x \<rightarrow>\<^sup>1 y == (x,y) \<in> contract"
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    36
| K:     "K\<bullet>x\<bullet>y \<rightarrow>\<^sup>1 x"
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    37
| S:     "S\<bullet>x\<bullet>y\<bullet>z \<rightarrow>\<^sup>1 (x\<bullet>z)\<bullet>(y\<bullet>z)"
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    38
| Ap1:   "x\<rightarrow>\<^sup>1y \<Longrightarrow> x\<bullet>z \<rightarrow>\<^sup>1 y\<bullet>z"
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    39
| Ap2:   "x\<rightarrow>\<^sup>1y \<Longrightarrow> z\<bullet>x \<rightarrow>\<^sup>1 z\<bullet>y"
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21210
diff changeset
    40
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21210
diff changeset
    41
abbreviation
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    42
  contract_rel :: "[comb,comb] \<Rightarrow> bool"   (infixl "\<rightarrow>" 50) where
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    43
  "x \<rightarrow> y == (x,y) \<in> contract\<^sup>*"
15816
4575c87dd25b x-symbol syntax
paulson
parents: 13075
diff changeset
    44
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    45
text \<open>
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    46
  Inductive definition of parallel contractions, \<open>\<Rrightarrow>\<^sup>1\<close> and
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    47
  (multi-step) parallel reductions, \<open>\<Rrightarrow>\<close>.
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    48
\<close>
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    49
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    50
inductive_set parcontract :: "(comb*comb) set"
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    51
  and parcontract_rel1 :: "[comb,comb] \<Rightarrow> bool"  (infixl "\<Rrightarrow>\<^sup>1" 50)
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    52
where
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    53
  "x \<Rrightarrow>\<^sup>1 y == (x,y) \<in> parcontract"
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    54
| refl:  "x \<Rrightarrow>\<^sup>1 x"
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    55
| K:     "K\<bullet>x\<bullet>y \<Rrightarrow>\<^sup>1 x"
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    56
| S:     "S\<bullet>x\<bullet>y\<bullet>z \<Rrightarrow>\<^sup>1 (x\<bullet>z)\<bullet>(y\<bullet>z)"
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    57
| Ap:    "[| x\<Rrightarrow>\<^sup>1y;  z\<Rrightarrow>\<^sup>1w |] ==> x\<bullet>z \<Rrightarrow>\<^sup>1 y\<bullet>w"
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21210
diff changeset
    58
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21210
diff changeset
    59
abbreviation
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    60
  parcontract_rel :: "[comb,comb] \<Rightarrow> bool"   (infixl "\<Rrightarrow>" 50) where
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    61
  "x \<Rrightarrow> y == (x,y) \<in> parcontract\<^sup>*"
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    62
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    63
text \<open>
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
    64
  Misc definitions.
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    65
\<close>
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    66
19736
wenzelm
parents: 16588
diff changeset
    67
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21210
diff changeset
    68
  I :: comb where
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    69
  "I = S\<bullet>K\<bullet>K"
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    70
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21210
diff changeset
    71
definition
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    72
  diamond   :: "('a * 'a)set \<Rightarrow> bool" where
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    73
    \<comment>\<open>confluence; Lambda/Commutation treats this more abstractly\<close>
19736
wenzelm
parents: 16588
diff changeset
    74
  "diamond(r) = (\<forall>x y. (x,y) \<in> r --> 
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
    75
                  (\<forall>y'. (x,y') \<in> r --> 
19736
wenzelm
parents: 16588
diff changeset
    76
                    (\<exists>z. (y,z) \<in> r & (y',z) \<in> r)))"
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
    77
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
    78
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    79
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
    80
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    81
text\<open>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
    82
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    83
text\<open>Strip lemma.  
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    84
   The induction hypothesis covers all but the last diamond of the strip.\<close>
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
    85
lemma diamond_strip_lemmaE [rule_format]: 
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    86
    "[| diamond(r);  (x,y) \<in> r\<^sup>* |] ==>   
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    87
          \<forall>y'. (x,y') \<in> r --> (\<exists>z. (y',z) \<in> r\<^sup>* & (y,z) \<in> r)"
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
    88
apply (unfold diamond_def)
16563
a92f96951355 meson method taking an argument list
paulson
parents: 16417
diff changeset
    89
apply (erule rtrancl_induct)
a92f96951355 meson method taking an argument list
paulson
parents: 16417
diff changeset
    90
apply (meson rtrancl_refl)
a92f96951355 meson method taking an argument list
paulson
parents: 16417
diff changeset
    91
apply (meson rtrancl_trans r_into_rtrancl)
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
    92
done
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
    93
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
    94
lemma diamond_rtrancl: "diamond(r) \<Longrightarrow> diamond(r\<^sup>*)"
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
    95
apply (simp (no_asm_simp) add: diamond_def)
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
    96
apply (rule impI [THEN allI, THEN allI])
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
    97
apply (erule rtrancl_induct, blast)
16588
8de758143786 stricter first-order check for meson
paulson
parents: 16563
diff changeset
    98
apply (meson rtrancl_trans r_into_rtrancl diamond_strip_lemmaE)
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
    99
done
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   100
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   101
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   102
subsection \<open>Non-contraction results\<close>
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   103
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   104
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
   105
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   106
inductive_cases
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
   107
      K_contractE [elim!]: "K \<rightarrow>\<^sup>1 r"
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
   108
  and S_contractE [elim!]: "S \<rightarrow>\<^sup>1 r"
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
   109
  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
   110
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   111
declare contract.K [intro!] contract.S [intro!]
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   112
declare contract.Ap1 [intro] contract.Ap2 [intro]
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   113
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
   114
lemma I_contract_E [elim!]: "I \<rightarrow>\<^sup>1 z \<Longrightarrow> P"
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   115
by (unfold I_def, blast)
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   116
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
   117
lemma K1_contractD [elim!]: "K\<bullet>x \<rightarrow>\<^sup>1 z \<Longrightarrow> (\<exists>x'. z = K\<bullet>x' & x \<rightarrow>\<^sup>1 x')"
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   118
by blast
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   119
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
   120
lemma Ap_reduce1 [intro]: "x \<rightarrow> y \<Longrightarrow> x\<bullet>z \<rightarrow> y\<bullet>z"
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   121
apply (erule rtrancl_induct)
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   122
apply (blast intro: rtrancl_trans)+
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   123
done
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   124
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
   125
lemma Ap_reduce2 [intro]: "x \<rightarrow> y \<Longrightarrow> z\<bullet>x \<rightarrow> z\<bullet>y"
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   126
apply (erule rtrancl_induct)
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   127
apply (blast intro: rtrancl_trans)+
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   128
done
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   129
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
   130
text \<open>Counterexample to the diamond property for @{term "x \<rightarrow>\<^sup>1 y"}\<close>
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   131
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   132
lemma not_diamond_contract: "~ diamond(contract)"
35621
1c084dda4c3c simplified
paulson
parents: 23746
diff changeset
   133
by (unfold diamond_def, metis S_contractE contract.K) 
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   134
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   135
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   136
subsection \<open>Results about Parallel Contraction\<close>
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   137
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   138
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
   139
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   140
inductive_cases
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
   141
      K_parcontractE [elim!]: "K \<Rrightarrow>\<^sup>1 r"
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
   142
  and S_parcontractE [elim!]: "S \<Rrightarrow>\<^sup>1 r"
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
   143
  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
   144
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   145
declare parcontract.intros [intro]
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   146
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   147
(*** Basic properties of parallel contraction ***)
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   148
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   149
subsection \<open>Basic properties of parallel contraction\<close>
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   150
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
   151
lemma K1_parcontractD [dest!]: "K\<bullet>x \<Rrightarrow>\<^sup>1 z \<Longrightarrow> (\<exists>x'. z = K\<bullet>x' & x \<Rrightarrow>\<^sup>1 x')"
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   152
by blast
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   153
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
   154
lemma S1_parcontractD [dest!]: "S\<bullet>x \<Rrightarrow>\<^sup>1 z \<Longrightarrow> (\<exists>x'. z = S\<bullet>x' & x \<Rrightarrow>\<^sup>1 x')"
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   155
by blast
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   156
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   157
lemma S2_parcontractD [dest!]:
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
   158
     "S\<bullet>x\<bullet>y \<Rrightarrow>\<^sup>1 z \<Longrightarrow> (\<exists>x' y'. z = S\<bullet>x'\<bullet>y' & x \<Rrightarrow>\<^sup>1 x' & y \<Rrightarrow>\<^sup>1 y')"
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   159
by blast
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   160
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   161
text\<open>The rules above 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
   162
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   163
text\<open>Church-Rosser property for parallel contraction\<close>
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   164
lemma diamond_parcontract: "diamond parcontract"
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   165
apply (unfold diamond_def)
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   166
apply (rule impI [THEN allI, THEN allI])
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   167
apply (erule parcontract.induct, fast+)
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   168
done
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   169
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   170
text \<open>
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
   171
  \<^medskip>
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
   172
  Equivalence of @{prop "p \<rightarrow> q"} and @{prop "p \<Rrightarrow> q"}.
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   173
\<close>
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   174
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
   175
lemma contract_subset_parcontract: "contract \<subseteq> parcontract"
35621
1c084dda4c3c simplified
paulson
parents: 23746
diff changeset
   176
by (auto, erule contract.induct, blast+)
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   177
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   178
text\<open>Reductions: simply throw together reflexivity, transitivity and
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   179
  the one-step reductions\<close>
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   180
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   181
declare r_into_rtrancl [intro]  rtrancl_trans [intro]
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   182
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   183
(*Example only: not used*)
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
   184
lemma reduce_I: "I\<bullet>x \<rightarrow> x"
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   185
by (unfold I_def, blast)
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   186
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
   187
lemma parcontract_subset_reduce: "parcontract \<subseteq> contract\<^sup>*"
35621
1c084dda4c3c simplified
paulson
parents: 23746
diff changeset
   188
by (auto, erule parcontract.induct, blast+)
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   189
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
   190
lemma reduce_eq_parreduce: "contract\<^sup>* = parcontract\<^sup>*"
35621
1c084dda4c3c simplified
paulson
parents: 23746
diff changeset
   191
by (metis contract_subset_parcontract parcontract_subset_reduce rtrancl_subset)
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   192
61993
89206877f0ee more symbols;
wenzelm
parents: 60530
diff changeset
   193
theorem diamond_reduce: "diamond(contract\<^sup>*)"
13075
d3e1d554cd6d conversion of some HOL/Induct proof scripts to Isar
paulson
parents: 11359
diff changeset
   194
by (simp add: reduce_eq_parreduce diamond_rtrancl diamond_parcontract)
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   195
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   196
end