src/HOL/ZF/LProd.thy
author avigad
Fri, 17 Jul 2009 13:12:18 -0400
changeset 32047 c141f139ce26
parent 29667 53103fc8ffa3
child 35416 d8d7d1b785af
permissions -rw-r--r--
Changed fact_Suc_nat back to fact_Suc
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
     1
(*  Title:      HOL/ZF/LProd.thy
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
     2
    ID:         $Id$
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
     3
    Author:     Steven Obua
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
     4
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
     5
    Introduces the lprod relation.
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
     6
    See "Partizan Games in Isabelle/HOLZF", available from http://www4.in.tum.de/~obua/partizan
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
     7
*)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
     8
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
     9
theory LProd 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    10
imports Multiset
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    11
begin
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    12
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
    13
inductive_set
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
    14
  lprod :: "('a * 'a) set \<Rightarrow> ('a list * 'a list) set"
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
    15
  for R :: "('a * 'a) set"
22282
71b4aefad227 - Adapted to new inductive definition package
berghofe
parents: 19769
diff changeset
    16
where
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
    17
  lprod_single[intro!]: "(a, b) \<in> R \<Longrightarrow> ([a], [b]) \<in> lprod R"
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
    18
| lprod_list[intro!]: "(ah@at, bh@bt) \<in> lprod R \<Longrightarrow> (a,b) \<in> R \<or> a = b \<Longrightarrow> (ah@a#at, bh@b#bt) \<in> lprod R"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    19
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
    20
lemma "(as,bs) \<in> lprod R \<Longrightarrow> length as = length bs"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    21
  apply (induct as bs rule: lprod.induct)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    22
  apply auto
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    23
  done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    24
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
    25
lemma "(as, bs) \<in> lprod R \<Longrightarrow> 1 \<le> length as \<and> 1 \<le> length bs"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    26
  apply (induct as bs rule: lprod.induct)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    27
  apply auto
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    28
  done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    29
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
    30
lemma lprod_subset_elem: "(as, bs) \<in> lprod S \<Longrightarrow> S \<subseteq> R \<Longrightarrow> (as, bs) \<in> lprod R"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    31
  apply (induct as bs rule: lprod.induct)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    32
  apply (auto)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    33
  done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    34
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
    35
lemma lprod_subset: "S \<subseteq> R \<Longrightarrow> lprod S \<subseteq> lprod R"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    36
  by (auto intro: lprod_subset_elem)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    37
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
    38
lemma lprod_implies_mult: "(as, bs) \<in> lprod R \<Longrightarrow> trans R \<Longrightarrow> (multiset_of as, multiset_of bs) \<in> mult R"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    39
proof (induct as bs rule: lprod.induct)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    40
  case (lprod_single a b)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    41
  note step = one_step_implies_mult[
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    42
    where r=R and I="{#}" and K="{#a#}" and J="{#b#}", simplified]    
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    43
  show ?case by (auto intro: lprod_single step)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    44
next
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
    45
  case (lprod_list ah at bh bt a b)
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
    46
  from prems have transR: "trans R" by auto
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    47
  have as: "multiset_of (ah @ a # at) = multiset_of (ah @ at) + {#a#}" (is "_ = ?ma + _")
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 23771
diff changeset
    48
    by (simp add: algebra_simps)
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    49
  have bs: "multiset_of (bh @ b # bt) = multiset_of (bh @ bt) + {#b#}" (is "_ = ?mb + _")
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 23771
diff changeset
    50
    by (simp add: algebra_simps)
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
    51
  from prems have "(?ma, ?mb) \<in> mult R"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    52
    by auto
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    53
  with mult_implies_one_step[OF transR] have 
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
    54
    "\<exists>I J K. ?mb = I + J \<and> ?ma = I + K \<and> J \<noteq> {#} \<and> (\<forall>k\<in>set_of K. \<exists>j\<in>set_of J. (k, j) \<in> R)"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    55
    by blast
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    56
  then obtain I J K where 
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
    57
    decomposed: "?mb = I + J \<and> ?ma = I + K \<and> J \<noteq> {#} \<and> (\<forall>k\<in>set_of K. \<exists>j\<in>set_of J. (k, j) \<in> R)"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    58
    by blast   
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    59
  show ?case
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    60
  proof (cases "a = b")
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    61
    case True    
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
    62
    have "((I + {#b#}) + K, (I + {#b#}) + J) \<in> mult R"
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
    63
      apply (rule one_step_implies_mult[OF transR])
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    64
      apply (auto simp add: decomposed)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    65
      done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    66
    then show ?thesis
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    67
      apply (simp only: as bs)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    68
      apply (simp only: decomposed True)
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 23771
diff changeset
    69
      apply (simp add: algebra_simps)
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    70
      done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    71
  next
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    72
    case False
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
    73
    from False lprod_list have False: "(a, b) \<in> R" by blast
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
    74
    have "(I + (K + {#a#}), I + (J + {#b#})) \<in> mult R"
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
    75
      apply (rule one_step_implies_mult[OF transR])
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    76
      apply (auto simp add: False decomposed)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    77
      done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    78
    then show ?thesis
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    79
      apply (simp only: as bs)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    80
      apply (simp only: decomposed)
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 23771
diff changeset
    81
      apply (simp add: algebra_simps)
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    82
      done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    83
  qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    84
qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    85
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    86
lemma wf_lprod[recdef_wf,simp,intro]:
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
    87
  assumes wf_R: "wf R"
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
    88
  shows "wf (lprod R)"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    89
proof -
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
    90
  have subset: "lprod (R^+) \<subseteq> inv_image (mult (R^+)) multiset_of"
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
    91
    by (auto simp add: lprod_implies_mult trans_trancl)
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
    92
  note lprodtrancl = wf_subset[OF wf_inv_image[where r="mult (R^+)" and f="multiset_of", 
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
    93
    OF wf_mult[OF wf_trancl[OF wf_R]]], OF subset]
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
    94
  note lprod = wf_subset[OF lprodtrancl, where p="lprod R", OF lprod_subset, simplified]
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    95
  show ?thesis by (auto intro: lprod)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    96
qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    97
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    98
constdefs
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
    99
  gprod_2_2 :: "('a * 'a) set \<Rightarrow> (('a * 'a) * ('a * 'a)) set"
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
   100
  "gprod_2_2 R \<equiv> { ((a,b), (c,d)) . (a = c \<and> (b,d) \<in> R) \<or> (b = d \<and> (a,c) \<in> R) }"
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
   101
  gprod_2_1 :: "('a * 'a) set \<Rightarrow> (('a * 'a) * ('a * 'a)) set"
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
   102
  "gprod_2_1 R \<equiv>  { ((a,b), (c,d)) . (a = d \<and> (b,c) \<in> R) \<or> (b = c \<and> (a,d) \<in> R) }"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   103
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
   104
lemma lprod_2_3: "(a, b) \<in> R \<Longrightarrow> ([a, c], [b, c]) \<in> lprod R"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   105
  by (auto intro: lprod_list[where a=c and b=c and 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   106
    ah = "[a]" and at = "[]" and bh="[b]" and bt="[]", simplified]) 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   107
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
   108
lemma lprod_2_4: "(a, b) \<in> R \<Longrightarrow> ([c, a], [c, b]) \<in> lprod R"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   109
  by (auto intro: lprod_list[where a=c and b=c and 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   110
    ah = "[]" and at = "[a]" and bh="[]" and bt="[b]", simplified])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   111
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
   112
lemma lprod_2_1: "(a, b) \<in> R \<Longrightarrow> ([c, a], [b, c]) \<in> lprod R"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   113
  by (auto intro: lprod_list[where a=c and b=c and 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   114
    ah = "[]" and at = "[a]" and bh="[b]" and bt="[]", simplified]) 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   115
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
   116
lemma lprod_2_2: "(a, b) \<in> R \<Longrightarrow> ([a, c], [c, b]) \<in> lprod R"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   117
  by (auto intro: lprod_list[where a=c and b=c and 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   118
    ah = "[a]" and at = "[]" and bh="[]" and bt="[b]", simplified])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   119
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   120
lemma [recdef_wf, simp, intro]: 
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
   121
  assumes wfR: "wf R" shows "wf (gprod_2_1 R)"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   122
proof -
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
   123
  have "gprod_2_1 R \<subseteq> inv_image (lprod R) (\<lambda> (a,b). [a,b])"
19769
c40ce2de2020 Added [simp]-lemmas "in_inv_image" and "in_lex_prod" in the spirit of "in_measure".
krauss
parents: 19203
diff changeset
   124
    by (auto simp add: gprod_2_1_def lprod_2_1 lprod_2_2)
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   125
  with wfR show ?thesis
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
   126
    by (rule_tac wf_subset, auto)
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   127
qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   128
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   129
lemma [recdef_wf, simp, intro]: 
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
   130
  assumes wfR: "wf R" shows "wf (gprod_2_2 R)"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   131
proof -
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
   132
  have "gprod_2_2 R \<subseteq> inv_image (lprod R) (\<lambda> (a,b). [a,b])"
19769
c40ce2de2020 Added [simp]-lemmas "in_inv_image" and "in_lex_prod" in the spirit of "in_measure".
krauss
parents: 19203
diff changeset
   133
    by (auto simp add: gprod_2_2_def lprod_2_3 lprod_2_4)
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   134
  with wfR show ?thesis
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
   135
    by (rule_tac wf_subset, auto)
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   136
qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   137
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
   138
lemma lprod_3_1: assumes "(x', x) \<in> R" shows "([y, z, x'], [x, y, z]) \<in> lprod R"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   139
  apply (rule lprod_list[where a="y" and b="y" and ah="[]" and at="[z,x']" and bh="[x]" and bt="[z]", simplified])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   140
  apply (auto simp add: lprod_2_1 prems)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   141
  done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   142
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
   143
lemma lprod_3_2: assumes "(z',z) \<in> R" shows "([z', x, y], [x,y,z]) \<in> lprod R"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   144
  apply (rule lprod_list[where a="y" and b="y" and ah="[z',x]" and at="[]" and bh="[x]" and bt="[z]", simplified])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   145
  apply (auto simp add: lprod_2_2 prems)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   146
  done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   147
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
   148
lemma lprod_3_3: assumes xr: "(xr, x) \<in> R" shows "([xr, y, z], [x, y, z]) \<in> lprod R"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   149
  apply (rule lprod_list[where a="y" and b="y" and ah="[xr]" and at="[z]" and bh="[x]" and bt="[z]", simplified])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   150
  apply (simp add: xr lprod_2_3)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   151
  done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   152
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
   153
lemma lprod_3_4: assumes yr: "(yr, y) \<in> R" shows "([x, yr, z], [x, y, z]) \<in> lprod R"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   154
  apply (rule lprod_list[where a="x" and b="x" and ah="[]" and at="[yr,z]" and bh="[]" and bt="[y,z]", simplified])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   155
  apply (simp add: yr lprod_2_3)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   156
  done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   157
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
   158
lemma lprod_3_5: assumes zr: "(zr, z) \<in> R" shows "([x, y, zr], [x, y, z]) \<in> lprod R"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   159
  apply (rule lprod_list[where a="x" and b="x" and ah="[]" and at="[y,zr]" and bh="[]" and bt="[y,z]", simplified])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   160
  apply (simp add: zr lprod_2_4)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   161
  done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   162
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
   163
lemma lprod_3_6: assumes y': "(y', y) \<in> R" shows "([x, z, y'], [x, y, z]) \<in> lprod R"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   164
  apply (rule lprod_list[where a="z" and b="z" and ah="[x]" and at="[y']" and bh="[x,y]" and bt="[]", simplified])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   165
  apply (simp add: y' lprod_2_4)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   166
  done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   167
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
   168
lemma lprod_3_7: assumes z': "(z',z) \<in> R" shows "([x, z', y], [x, y, z]) \<in> lprod R"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   169
  apply (rule lprod_list[where a="y" and b="y" and ah="[x, z']" and at="[]" and bh="[x]" and bt="[z]", simplified])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   170
  apply (simp add: z' lprod_2_4)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   171
  done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   172
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   173
constdefs
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   174
   perm :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a set \<Rightarrow> bool"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   175
   "perm f A \<equiv> inj_on f A \<and> f ` A = A"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   176
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
   177
lemma "((as,bs) \<in> lprod R) = 
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   178
  (\<exists> f. perm f {0 ..< (length as)} \<and> 
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
   179
  (\<forall> j. j < length as \<longrightarrow> ((nth as j, nth bs (f j)) \<in> R \<or> (nth as j = nth bs (f j)))) \<and> 
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
   180
  (\<exists> i. i < length as \<and> (nth as i, nth bs (f i)) \<in> R))"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   181
oops
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   182
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 23477
diff changeset
   183
lemma "trans R \<Longrightarrow> (ah@a#at, bh@b#bt) \<in> lprod R \<Longrightarrow> (b, a) \<in> R \<or> a = b \<Longrightarrow> (ah@at, bh@bt) \<in> lprod R" 
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   184
oops
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   185
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   186
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   187
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   188
end