QuotScript.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Wed, 28 Oct 2009 14:59:24 +0100
changeset 217 9cc87d02190a
parent 188 b8485573548d
child 228 268a727b0f10
permissions -rw-r--r--
First experiments with Lambda
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
0
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory QuotScript
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
imports Main
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
begin
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
definition 
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
  "EQUIV E \<equiv> \<forall>x y. E x y = (E x = E y)" 
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
definition
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
  "REFL E \<equiv> \<forall>x. E x x"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
definition 
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
  "SYM E \<equiv> \<forall>x y. E x y \<longrightarrow> E y x"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
definition
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
  "TRANS E \<equiv> \<forall>x y z. E x y \<and> E y z \<longrightarrow> E x z"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
lemma EQUIV_REFL_SYM_TRANS:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
  shows "EQUIV E = (REFL E \<and> SYM E \<and> TRANS E)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
unfolding EQUIV_def REFL_def SYM_def TRANS_def expand_fun_eq
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
by (blast)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
217
9cc87d02190a First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 188
diff changeset
    22
lemma EQUIV_REFL:
9cc87d02190a First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 188
diff changeset
    23
  shows "EQUIV E ==> REFL E"
9cc87d02190a First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 188
diff changeset
    24
  by (simp add: EQUIV_REFL_SYM_TRANS)
9cc87d02190a First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 188
diff changeset
    25
0
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
definition
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
  "PART_EQUIV E \<equiv> (\<exists>x. E x x) \<and> (\<forall>x y. E x y = (E x x \<and> E y y \<and> (E x = E y)))"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
lemma EQUIV_IMP_PART_EQUIV:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
  assumes a: "EQUIV E"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
  shows "PART_EQUIV E"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
using a unfolding EQUIV_def PART_EQUIV_def
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
by auto
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
definition
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
  "QUOTIENT E Abs Rep \<equiv> (\<forall>a. Abs (Rep a) = a) \<and> 
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
                        (\<forall>a. E (Rep a) (Rep a)) \<and> 
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
                        (\<forall>r s. E r s = (E r r \<and> E s s \<and> (Abs r = Abs s)))"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
lemma QUOTIENT_ABS_REP:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
  assumes a: "QUOTIENT E Abs Rep"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
  shows "Abs (Rep a) = a" 
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
using a unfolding QUOTIENT_def
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
by simp
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
lemma QUOTIENT_REP_REFL:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
  assumes a: "QUOTIENT E Abs Rep"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
  shows "E (Rep a) (Rep a)" 
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
using a unfolding QUOTIENT_def
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
by blast
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
lemma QUOTIENT_REL:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
  assumes a: "QUOTIENT E Abs Rep"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
  shows " E r s = (E r r \<and> E s s \<and> (Abs r = Abs s))"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
using a unfolding QUOTIENT_def
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
by blast
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
lemma QUOTIENT_REL_ABS:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
  assumes a: "QUOTIENT E Abs Rep"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
  shows "E r s \<Longrightarrow> Abs r = Abs s"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
using a unfolding QUOTIENT_def
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
by blast
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
lemma QUOTIENT_REL_ABS_EQ:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
  assumes a: "QUOTIENT E Abs Rep"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
  shows "E r r \<Longrightarrow> E s s \<Longrightarrow> E r s = (Abs r = Abs s)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
using a unfolding QUOTIENT_def
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
by blast
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    69
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
lemma QUOTIENT_REL_REP:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    71
  assumes a: "QUOTIENT E Abs Rep"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    72
  shows "E (Rep a) (Rep b) = (a = b)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
using a unfolding QUOTIENT_def
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
by metis
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    76
lemma QUOTIENT_REP_ABS:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    77
  assumes a: "QUOTIENT E Abs Rep"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
  shows "E r r \<Longrightarrow> E (Rep (Abs r)) r"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79
using a unfolding QUOTIENT_def
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80
by blast
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
lemma IDENTITY_EQUIV:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
  shows "EQUIV (op =)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    84
unfolding EQUIV_def
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
by auto
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
lemma IDENTITY_QUOTIENT:
126
9cb8f9a59402 Partial simplification of the proof
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 113
diff changeset
    88
  shows "QUOTIENT (op =) id id"
9cb8f9a59402 Partial simplification of the proof
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 113
diff changeset
    89
unfolding QUOTIENT_def id_def
0
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
by blast
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
lemma QUOTIENT_SYM:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
  assumes a: "QUOTIENT E Abs Rep"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
  shows "SYM E"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    95
using a unfolding QUOTIENT_def SYM_def
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    96
by metis
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
lemma QUOTIENT_TRANS:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
  assumes a: "QUOTIENT E Abs Rep"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
  shows "TRANS E"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   101
using a unfolding QUOTIENT_def TRANS_def
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
by metis
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
fun
93
ec29be471518 Manually regularized list_induct2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 0
diff changeset
   105
  prod_rel
ec29be471518 Manually regularized list_induct2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 0
diff changeset
   106
where
ec29be471518 Manually regularized list_induct2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 0
diff changeset
   107
  "prod_rel r1 r2 = (\<lambda>(a,b) (c,d). r1 a c \<and> r2 b d)"
ec29be471518 Manually regularized list_induct2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 0
diff changeset
   108
ec29be471518 Manually regularized list_induct2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 0
diff changeset
   109
fun
112
0d6d37d0589d Progressing with the proof
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 96
diff changeset
   110
  fun_map
0
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
where
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
  "fun_map f g h x = g (h (f x))"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
112
0d6d37d0589d Progressing with the proof
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 96
diff changeset
   114
0
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
abbreviation
112
0d6d37d0589d Progressing with the proof
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 96
diff changeset
   116
  fun_map_syn (infixr "--->" 55)
0
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
where
112
0d6d37d0589d Progressing with the proof
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 96
diff changeset
   118
  "f ---> g \<equiv> fun_map f g"
0
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
lemma FUN_MAP_I:
126
9cb8f9a59402 Partial simplification of the proof
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 113
diff changeset
   121
  shows "(id ---> id) = id"
9cb8f9a59402 Partial simplification of the proof
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 113
diff changeset
   122
by (simp add: expand_fun_eq id_def)
0
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
lemma IN_FUN:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
  shows "x \<in> ((f ---> g) s) = g (f x \<in> s)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
by (simp add: mem_def)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
fun
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
  FUN_REL 
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
where
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
  "FUN_REL E1 E2 f g = (\<forall>x y. E1 x y \<longrightarrow> E2 (f x) (g y))"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
abbreviation
173
7cf227756e2a Finally completely lift the previously lifted theorems + clean some old stuff
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 172
diff changeset
   134
  FUN_REL_syn ("_ ===> _")
0
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   135
where
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   136
  "E1 ===> E2 \<equiv> FUN_REL E1 E2"  
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   137
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   138
lemma FUN_REL_EQ:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   139
  "(op =) ===> (op =) = (op =)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   140
by (simp add: expand_fun_eq)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   141
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   142
lemma FUN_QUOTIENT:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   143
  assumes q1: "QUOTIENT R1 abs1 rep1"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   144
  and     q2: "QUOTIENT R2 abs2 rep2"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   145
  shows "QUOTIENT (R1 ===> R2) (rep1 ---> abs2) (abs1 ---> rep2)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   146
proof -
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   147
  have "\<forall>a. (rep1 ---> abs2) ((abs1 ---> rep2) a) = a"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   148
    apply(simp add: expand_fun_eq)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
    using q1 q2
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   150
    apply(simp add: QUOTIENT_def)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   151
    done
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   152
  moreover
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   153
  have "\<forall>a. (R1 ===> R2) ((abs1 ---> rep2) a) ((abs1 ---> rep2) a)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
    apply(auto)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
    using q1 q2 unfolding QUOTIENT_def
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   156
    apply(metis)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
    done
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
  moreover
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
  have "\<forall>r s. (R1 ===> R2) r s = ((R1 ===> R2) r r \<and> (R1 ===> R2) s s \<and> 
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   160
        (rep1 ---> abs2) r  = (rep1 ---> abs2) s)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   161
    apply(auto simp add: expand_fun_eq)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   162
    using q1 q2 unfolding QUOTIENT_def
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   163
    apply(metis)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   164
    using q1 q2 unfolding QUOTIENT_def
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   165
    apply(metis)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   166
    using q1 q2 unfolding QUOTIENT_def
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   167
    apply(metis)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   168
    using q1 q2 unfolding QUOTIENT_def
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   169
    apply(metis)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   170
    done
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   171
  ultimately
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
  show "QUOTIENT (R1 ===> R2) (rep1 ---> abs2) (abs1 ---> rep2)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   173
    unfolding QUOTIENT_def by blast
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
qed
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
definition
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
  Respects
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   178
where
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   179
  "Respects R x \<equiv> (R x x)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   180
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   181
lemma IN_RESPECTS:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
  shows "(x \<in> Respects R) = R x x"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   183
unfolding mem_def Respects_def by simp
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
lemma RESPECTS_THM:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   186
  shows "Respects (R1 ===> R2) f = (\<forall>x y. R1 x y \<longrightarrow> R2 (f x) (f y))"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   187
unfolding Respects_def
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   188
by (simp add: expand_fun_eq) 
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   190
lemma RESPECTS_MP:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   191
  assumes a: "Respects (R1 ===> R2) f"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   192
  and     b: "R1 x y"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   193
  shows "R2 (f x) (f y)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   194
using a b unfolding Respects_def
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
by simp
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   196
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   197
lemma RESPECTS_REP_ABS:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   198
  assumes a: "QUOTIENT R1 Abs1 Rep1"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   199
  and     b: "Respects (R1 ===> R2) f"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   200
  and     c: "R1 x x"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
  shows "R2 (f (Rep1 (Abs1 x))) (f x)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   202
using a b[simplified RESPECTS_THM] c unfolding QUOTIENT_def
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   203
by blast
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   204
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   205
lemma RESPECTS_o:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
  assumes a: "Respects (R2 ===> R3) f"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   207
  and     b: "Respects (R1 ===> R2) g"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   208
  shows "Respects (R1 ===> R3) (f o g)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   209
using a b unfolding Respects_def
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   210
by simp
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   211
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   212
(*
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   213
definition
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   214
  "RES_EXISTS_EQUIV R P \<equiv> (\<exists>x \<in> Respects R. P x) \<and> 
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   215
                          (\<forall>x\<in> Respects R. \<forall>y\<in> Respects R. P x \<and> P y \<longrightarrow> R x y)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   216
*)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   217
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   218
lemma FUN_REL_EQ_REL:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   219
  assumes q1: "QUOTIENT R1 Abs1 Rep1"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   220
  and     q2: "QUOTIENT R2 Abs2 Rep2"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   221
  shows "(R1 ===> R2) f g = ((Respects (R1 ===> R2) f) \<and> (Respects (R1 ===> R2) g) 
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   222
                             \<and> ((Rep1 ---> Abs2) f = (Rep1 ---> Abs2) g))"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   223
using FUN_QUOTIENT[OF q1 q2] unfolding Respects_def QUOTIENT_def expand_fun_eq
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   224
by blast
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   225
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   226
(* q1 and q2 not used; see next lemma *)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   227
lemma FUN_REL_MP:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   228
  assumes q1: "QUOTIENT R1 Abs1 Rep1"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   229
  and     q2: "QUOTIENT R2 Abs2 Rep2"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   230
  shows "(R1 ===> R2) f g \<Longrightarrow> R1 x y \<Longrightarrow> R2 (f x) (g y)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   231
by simp
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   232
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   233
lemma FUN_REL_IMP:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   234
  shows "(R1 ===> R2) f g \<Longrightarrow> R1 x y \<Longrightarrow> R2 (f x) (g y)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   235
by simp
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   236
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   237
lemma FUN_REL_EQUALS:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   238
  assumes q1: "QUOTIENT R1 Abs1 Rep1"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   239
  and     q2: "QUOTIENT R2 Abs2 Rep2"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   240
  and     r1: "Respects (R1 ===> R2) f"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   241
  and     r2: "Respects (R1 ===> R2) g" 
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   242
  shows "((Rep1 ---> Abs2) f = (Rep1 ---> Abs2) g) = (\<forall>x y. R1 x y \<longrightarrow> R2 (f x) (g y))"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   243
apply(rule_tac iffI)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
using FUN_QUOTIENT[OF q1 q2] r1 r2 unfolding QUOTIENT_def Respects_def
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   245
apply(metis FUN_REL_IMP)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   246
using r1 unfolding Respects_def expand_fun_eq
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   247
apply(simp (no_asm_use))
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   248
apply(metis QUOTIENT_REL[OF q2] QUOTIENT_REL_REP[OF q1])
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   249
done
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   250
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   251
(* ask Peter: FUN_REL_IMP used twice *) 
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   252
lemma FUN_REL_IMP2:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   253
  assumes q1: "QUOTIENT R1 Abs1 Rep1"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   254
  and     q2: "QUOTIENT R2 Abs2 Rep2"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   255
  and     r1: "Respects (R1 ===> R2) f"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   256
  and     r2: "Respects (R1 ===> R2) g" 
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   257
  and     a:  "(Rep1 ---> Abs2) f = (Rep1 ---> Abs2) g"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   258
  shows "R1 x y \<Longrightarrow> R2 (f x) (g y)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   259
using q1 q2 r1 r2 a
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   260
by (simp add: FUN_REL_EQUALS)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   261
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   262
lemma EQUALS_PRS:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   263
  assumes q: "QUOTIENT R Abs Rep"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   264
  shows "(x = y) = R (Rep x) (Rep y)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   265
by (simp add: QUOTIENT_REL_REP[OF q]) 
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   266
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   267
lemma EQUALS_RSP:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   268
  assumes q: "QUOTIENT R Abs Rep"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   269
  and     a: "R x1 x2" "R y1 y2"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   270
  shows "R x1 y1 = R x2 y2"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   271
using QUOTIENT_SYM[OF q] QUOTIENT_TRANS[OF q] unfolding SYM_def TRANS_def
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   272
using a by blast
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   273
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   274
lemma LAMBDA_PRS:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   275
  assumes q1: "QUOTIENT R1 Abs1 Rep1"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   276
  and     q2: "QUOTIENT R2 Abs2 Rep2"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   277
  shows "(\<lambda>x. f x) = (Rep1 ---> Abs2) (\<lambda>x. Rep2 (f (Abs1 x)))"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   278
unfolding expand_fun_eq
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   279
using QUOTIENT_ABS_REP[OF q1] QUOTIENT_ABS_REP[OF q2]
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   280
by simp
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   281
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   282
lemma LAMBDA_PRS1:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   283
  assumes q1: "QUOTIENT R1 Abs1 Rep1"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   284
  and     q2: "QUOTIENT R2 Abs2 Rep2"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   285
  shows "(\<lambda>x. f x) = (Rep1 ---> Abs2) (\<lambda>x. (Abs1 ---> Rep2) f x)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   286
unfolding expand_fun_eq
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   287
by (subst LAMBDA_PRS[OF q1 q2]) (simp)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   288
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   289
(* Ask Peter: assumption q1 and q2 not used and lemma is the 'identity' *)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   290
lemma LAMBDA_RSP:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   291
  assumes q1: "QUOTIENT R1 Abs1 Rep1"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   292
  and     q2: "QUOTIENT R2 Abs2 Rep2"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   293
  and     a: "(R1 ===> R2) f1 f2"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   294
  shows "(R1 ===> R2) (\<lambda>x. f1 x) (\<lambda>y. f2 y)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   295
by (rule a)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   296
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   297
(* ASK Peter about next four lemmas in quotientScript
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   298
lemma ABSTRACT_PRS:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   299
  assumes q1: "QUOTIENT R1 Abs1 Rep1"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   300
  and     q2: "QUOTIENT R2 Abs2 Rep2"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   301
  shows "f = (Rep1 ---> Abs2) ???"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   302
*)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   303
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   304
lemma LAMBDA_REP_ABS_RSP:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   305
  assumes r1: "\<And>r r'. R1 r r' \<Longrightarrow>R1 r (Rep1 (Abs1 r'))"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   306
  and     r2: "\<And>r r'. R2 r r' \<Longrightarrow>R2 r (Rep2 (Abs2 r'))"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   307
  shows "(R1 ===> R2) f1 f2 \<Longrightarrow> (R1 ===> R2) f1 ((Abs1 ---> Rep2) ((Rep1 ---> Abs2) f2))"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   308
using r1 r2 by auto
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   309
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   310
lemma REP_ABS_RSP:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   311
  assumes q: "QUOTIENT R Abs Rep"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   312
  and     a: "R x1 x2"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   313
  shows "R x1 (Rep (Abs x2))"
113
e3a963e6418b Symmetric version of REP_ABS_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 112
diff changeset
   314
  and   "R (Rep (Abs x1)) x2"
e3a963e6418b Symmetric version of REP_ABS_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 112
diff changeset
   315
proof -
e3a963e6418b Symmetric version of REP_ABS_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 112
diff changeset
   316
  show "R x1 (Rep (Abs x2))" 
e3a963e6418b Symmetric version of REP_ABS_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 112
diff changeset
   317
    using q a by (metis QUOTIENT_REL[OF q] QUOTIENT_ABS_REP[OF q] QUOTIENT_REP_REFL[OF q])
e3a963e6418b Symmetric version of REP_ABS_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 112
diff changeset
   318
next
e3a963e6418b Symmetric version of REP_ABS_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 112
diff changeset
   319
  show "R (Rep (Abs x1)) x2"
e3a963e6418b Symmetric version of REP_ABS_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 112
diff changeset
   320
    using q a by (metis QUOTIENT_REL[OF q] QUOTIENT_ABS_REP[OF q] QUOTIENT_REP_REFL[OF q] QUOTIENT_SYM[of q])
e3a963e6418b Symmetric version of REP_ABS_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 112
diff changeset
   321
qed
0
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   322
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   323
(* ----------------------------------------------------- *)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   324
(* Quantifiers: FORALL, EXISTS, EXISTS_UNIQUE,           *)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   325
(*              RES_FORALL, RES_EXISTS, RES_EXISTS_EQUIV *)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   326
(* ----------------------------------------------------- *)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   327
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   328
(* what is RES_FORALL *)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   329
(*--`!R (abs:'a -> 'b) rep. QUOTIENT R abs rep ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   330
         !f. $! f = RES_FORALL (respects R) ((abs --> I) f)`--*)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   331
(*as peter here *)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   332
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   333
(* bool theory: COND, LET *)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   334
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   335
lemma IF_PRS:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   336
  assumes q: "QUOTIENT R Abs Rep"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   337
  shows "If a b c = Abs (If a (Rep b) (Rep c))"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   338
using QUOTIENT_ABS_REP[OF q] by auto
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   339
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   340
(* ask peter: no use of q *)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   341
lemma IF_RSP:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   342
  assumes q: "QUOTIENT R Abs Rep"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   343
  and     a: "a1 = a2" "R b1 b2" "R c1 c2"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   344
  shows "R (If a1 b1 c1) (If a2 b2 c2)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   345
using a by auto
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   346
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   347
lemma LET_PRS:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   348
  assumes q1: "QUOTIENT R1 Abs1 Rep1"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   349
  and     q2: "QUOTIENT R2 Abs2 Rep2"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   350
  shows "Let x f = Abs2 (Let (Rep1 x) ((Abs1 ---> Rep2) f))"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   351
using QUOTIENT_ABS_REP[OF q1] QUOTIENT_ABS_REP[OF q2] by auto
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   352
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   353
lemma LET_RSP:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   354
  assumes q1: "QUOTIENT R1 Abs1 Rep1"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   355
  and     q2: "QUOTIENT R2 Abs2 Rep2"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   356
  and     a1: "(R1 ===> R2) f g"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   357
  and     a2: "R1 x y"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   358
  shows "R2 (Let x f) (Let y g)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   359
using FUN_REL_MP[OF q1 q2 a1] a2
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   360
by auto
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   361
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   362
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   363
(* ask peter what are literal_case *)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   364
(* literal_case_PRS *)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   365
(* literal_case_RSP *)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   366
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   367
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   368
(* FUNCTION APPLICATION *)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   369
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   370
lemma APPLY_PRS:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   371
  assumes q1: "QUOTIENT R1 Abs1 Rep1"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   372
  and     q2: "QUOTIENT R2 Abs2 Rep2"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   373
  shows "f x = Abs2 (((Abs1 ---> Rep2) f) (Rep1 x))"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   374
using QUOTIENT_ABS_REP[OF q1] QUOTIENT_ABS_REP[OF q2] by auto
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   375
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   376
(* ask peter: no use of q1 q2 *)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   377
lemma APPLY_RSP:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   378
  assumes q1: "QUOTIENT R1 Abs1 Rep1"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   379
  and     q2: "QUOTIENT R2 Abs2 Rep2"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   380
  and     a: "(R1 ===> R2) f g" "R1 x y"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   381
  shows "R2 (f x) (g y)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   382
using a by (rule FUN_REL_IMP)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   383
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   384
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   385
(* combinators: I, K, o, C, W *)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   386
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   387
lemma I_PRS:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   388
  assumes q: "QUOTIENT R Abs Rep"
126
9cb8f9a59402 Partial simplification of the proof
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 113
diff changeset
   389
  shows "id e = Abs (id (Rep e))"
0
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   390
using QUOTIENT_ABS_REP[OF q] by auto
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   391
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   392
lemma I_RSP:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   393
  assumes q: "QUOTIENT R Abs Rep"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   394
  and     a: "R e1 e2"
126
9cb8f9a59402 Partial simplification of the proof
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 113
diff changeset
   395
  shows "R (id e1) (id e2)"
0
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   396
using a by auto
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   397
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   398
lemma o_PRS:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   399
  assumes q1: "QUOTIENT R1 Abs1 Rep1"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   400
  and     q2: "QUOTIENT R2 Abs2 Rep2"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   401
  and     q3: "QUOTIENT R3 Abs3 Rep3"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   402
  shows "f o g = (Rep1 ---> Abs3) (((Abs2 ---> Rep3) f) o ((Abs1 ---> Rep2) g))"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   403
using QUOTIENT_ABS_REP[OF q1] QUOTIENT_ABS_REP[OF q2] QUOTIENT_ABS_REP[OF q3]
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   404
unfolding o_def expand_fun_eq
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   405
by simp
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   406
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   407
lemma o_RSP:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   408
  assumes q1: "QUOTIENT R1 Abs1 Rep1"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   409
  and     q2: "QUOTIENT R2 Abs2 Rep2"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   410
  and     q3: "QUOTIENT R3 Abs3 Rep3"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   411
  and     a1: "(R2 ===> R3) f1 f2"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   412
  and     a2: "(R1 ===> R2) g1 g2"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   413
  shows "(R1 ===> R3) (f1 o g1) (f2 o g2)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   414
using a1 a2 unfolding o_def expand_fun_eq
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   415
by (auto)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   416
96
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   417
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   418
(* TODO: Put the following lemmas in proper places *)
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   419
93
ec29be471518 Manually regularized list_induct2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 0
diff changeset
   420
lemma equiv_res_forall:
ec29be471518 Manually regularized list_induct2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 0
diff changeset
   421
  fixes P :: "'a \<Rightarrow> bool"
ec29be471518 Manually regularized list_induct2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 0
diff changeset
   422
  assumes a: "EQUIV E"
ec29be471518 Manually regularized list_induct2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 0
diff changeset
   423
  shows "Ball (Respects E) P = (All P)"
ec29be471518 Manually regularized list_induct2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 0
diff changeset
   424
  using a by (metis EQUIV_def IN_RESPECTS a)
ec29be471518 Manually regularized list_induct2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 0
diff changeset
   425
ec29be471518 Manually regularized list_induct2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 0
diff changeset
   426
lemma equiv_res_exists:
ec29be471518 Manually regularized list_induct2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 0
diff changeset
   427
  fixes P :: "'a \<Rightarrow> bool"
ec29be471518 Manually regularized list_induct2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 0
diff changeset
   428
  assumes a: "EQUIV E"
ec29be471518 Manually regularized list_induct2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 0
diff changeset
   429
  shows "Bex (Respects E) P = (Ex P)"
ec29be471518 Manually regularized list_induct2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 0
diff changeset
   430
  using a by (metis EQUIV_def IN_RESPECTS a)
ec29be471518 Manually regularized list_induct2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 0
diff changeset
   431
96
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   432
lemma FORALL_REGULAR:
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   433
  assumes a: "!x :: 'a. (P x --> Q x)"
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   434
  and     b: "All P"
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   435
  shows "All Q"
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   436
  using a b by (metis)
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   437
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   438
lemma EXISTS_REGULAR:
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   439
  assumes a: "!x :: 'a. (P x --> Q x)"
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   440
  and     b: "Ex P"
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   441
  shows "Ex Q"
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   442
  using a b by (metis)
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   443
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   444
lemma RES_FORALL_REGULAR:
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   445
  assumes a: "!x :: 'a. (R x --> P x --> Q x)"
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   446
  and     b: "Ball R P"
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   447
  shows "Ball R Q"
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   448
  using a b by (metis COMBC_def Collect_def Collect_mem_eq)
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   449
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   450
lemma RES_EXISTS_REGULAR:
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   451
  assumes a: "!x :: 'a. (R x --> P x --> Q x)"
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   452
  and     b: "Bex R P"
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   453
  shows "Bex R Q"
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   454
  using a b by (metis COMBC_def Collect_def Collect_mem_eq)
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   455
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   456
lemma LEFT_RES_FORALL_REGULAR:
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   457
  assumes a: "!x. (R x \<and> (Q x --> P x))"
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   458
  shows "Ball R Q ==> All P"
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   459
  using a
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   460
  apply (metis COMBC_def Collect_def Collect_mem_eq a)
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   461
  done
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   462
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   463
lemma RIGHT_RES_FORALL_REGULAR:
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   464
  assumes a: "!x :: 'a. (R x --> P x --> Q x)"
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   465
  shows "All P ==> Ball R Q"
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   466
  using a
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   467
  apply (metis COMBC_def Collect_def Collect_mem_eq a)
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   468
  done
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   469
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   470
lemma LEFT_RES_EXISTS_REGULAR:
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   471
  assumes a: "!x :: 'a. (R x --> Q x --> P x)"
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   472
  shows "Bex R Q ==> Ex P"
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   473
  using a by (metis COMBC_def Collect_def Collect_mem_eq)
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   474
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   475
lemma RIGHT_RES_EXISTS_REGULAR:
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   476
  assumes a: "!x :: 'a. (R x \<and> (P x --> Q x))"
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   477
  shows "Ex P \<Longrightarrow> Bex R Q"
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   478
  using a by (metis COMBC_def Collect_def Collect_mem_eq)
4da714704611 A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 95
diff changeset
   479
162
20f0b148cfe2 eqsubst_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 155
diff changeset
   480
(* TODO: Add similar *)
153
0288dd5b7ed4 The problems with 'abs' term.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 126
diff changeset
   481
lemma RES_FORALL_RSP:
0288dd5b7ed4 The problems with 'abs' term.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 126
diff changeset
   482
  shows "\<And>f g. (R ===> (op =)) f g ==>
0288dd5b7ed4 The problems with 'abs' term.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 126
diff changeset
   483
        (Ball (Respects R) f = Ball (Respects R) g)"
155
8b3d4806ad79 Removed an assumption
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 154
diff changeset
   484
  apply (simp add: FUN_REL.simps Ball_def IN_RESPECTS)
8b3d4806ad79 Removed an assumption
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 154
diff changeset
   485
  done
153
0288dd5b7ed4 The problems with 'abs' term.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 126
diff changeset
   486
171
13aab4c59096 More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 166
diff changeset
   487
lemma RES_EXISTS_RSP:
13aab4c59096 More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 166
diff changeset
   488
  shows "\<And>f g. (R ===> (op =)) f g ==>
13aab4c59096 More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 166
diff changeset
   489
        (Bex (Respects R) f = Bex (Respects R) g)"
13aab4c59096 More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 166
diff changeset
   490
  apply (simp add: FUN_REL.simps Bex_def IN_RESPECTS)
13aab4c59096 More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 166
diff changeset
   491
  done
13aab4c59096 More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 166
diff changeset
   492
188
b8485573548d Finished COND_PRS proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 187
diff changeset
   493
162
20f0b148cfe2 eqsubst_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 155
diff changeset
   494
lemma FORALL_PRS:
20f0b148cfe2 eqsubst_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 155
diff changeset
   495
  assumes a: "QUOTIENT R absf repf"
183
6acf9e001038 proved the two lemmas in QuotScript (reformulated them without leading forall)
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   496
  shows "All f = Ball (Respects R) ((absf ---> id) f)"
6acf9e001038 proved the two lemmas in QuotScript (reformulated them without leading forall)
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   497
  using a
6acf9e001038 proved the two lemmas in QuotScript (reformulated them without leading forall)
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   498
  unfolding QUOTIENT_def
6acf9e001038 proved the two lemmas in QuotScript (reformulated them without leading forall)
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   499
  by (metis IN_RESPECTS fun_map.simps id_apply)
162
20f0b148cfe2 eqsubst_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 155
diff changeset
   500
171
13aab4c59096 More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 166
diff changeset
   501
lemma EXISTS_PRS:
13aab4c59096 More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 166
diff changeset
   502
  assumes a: "QUOTIENT R absf repf"
183
6acf9e001038 proved the two lemmas in QuotScript (reformulated them without leading forall)
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   503
  shows "Ex f = Bex (Respects R) ((absf ---> id) f)"
6acf9e001038 proved the two lemmas in QuotScript (reformulated them without leading forall)
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   504
  using a
6acf9e001038 proved the two lemmas in QuotScript (reformulated them without leading forall)
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   505
  unfolding QUOTIENT_def
6acf9e001038 proved the two lemmas in QuotScript (reformulated them without leading forall)
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   506
  by (metis COMBC_def Collect_def Collect_mem_eq IN_RESPECTS fun_map.simps id_apply mem_def)
187
f8fc085db38f Cleaning and fixing.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 183
diff changeset
   507
f8fc085db38f Cleaning and fixing.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 183
diff changeset
   508
lemma COND_PRS:
f8fc085db38f Cleaning and fixing.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 183
diff changeset
   509
  assumes a: "QUOTIENT R absf repf"
f8fc085db38f Cleaning and fixing.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 183
diff changeset
   510
  shows "(if a then b else c) = absf (if a then repf b else repf c)"
188
b8485573548d Finished COND_PRS proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 187
diff changeset
   511
  using a unfolding QUOTIENT_def by auto
187
f8fc085db38f Cleaning and fixing.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 183
diff changeset
   512
f8fc085db38f Cleaning and fixing.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 183
diff changeset
   513
(* These are the fixed versions, general ones need to be proved. *)
f8fc085db38f Cleaning and fixing.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 183
diff changeset
   514
lemma ho_all_prs:
183
6acf9e001038 proved the two lemmas in QuotScript (reformulated them without leading forall)
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   515
  shows "op = ===> op = ===> op = All All"
6acf9e001038 proved the two lemmas in QuotScript (reformulated them without leading forall)
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   516
  by auto
171
13aab4c59096 More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 166
diff changeset
   517
183
6acf9e001038 proved the two lemmas in QuotScript (reformulated them without leading forall)
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   518
lemma ho_ex_prs: 
6acf9e001038 proved the two lemmas in QuotScript (reformulated them without leading forall)
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   519
  shows "op = ===> op = ===> op = Ex Ex"
6acf9e001038 proved the two lemmas in QuotScript (reformulated them without leading forall)
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   520
  by auto
171
13aab4c59096 More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 166
diff changeset
   521
93
ec29be471518 Manually regularized list_induct2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 0
diff changeset
   522
end
95
8c3a35da4560 Proving the proper RepAbs version
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 93
diff changeset
   523