LamEx.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Tue, 03 Nov 2009 14:04:21 +0100
changeset 265 5f3b364d4765
parent 259 22c199522bef
child 267 3764566c1151
permissions -rw-r--r--
Preparing infrastructure for general FORALL_PRS
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
201
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory LamEx
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
imports Nominal QuotMain
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
begin
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
atom_decl name
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
243
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
     7
thm abs_fresh(1)
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
     8
201
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
nominal_datatype rlam =
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
  rVar "name"
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
| rApp "rlam" "rlam"
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
| rLam "name" "rlam"
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
259
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
    14
print_theorems
201
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
243
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    16
function
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    17
  rfv :: "rlam \<Rightarrow> name set"
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    18
where
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    19
  rfv_var: "rfv (rVar a) = {a}"
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    20
| rfv_app: "rfv (rApp t1 t2) = (rfv t1) \<union> (rfv t2)"
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    21
| rfv_lam: "rfv (rLam a t) = (rfv t) - {a}"
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    22
sorry
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    23
247
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
    24
termination rfv sorry
243
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    25
246
6da7d538e5e0 changed the order of rfv and reformulated a3 with rfv
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
    26
inductive 
6da7d538e5e0 changed the order of rfv and reformulated a3 with rfv
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
    27
  alpha :: "rlam \<Rightarrow> rlam \<Rightarrow> bool" ("_ \<approx> _" [100, 100] 100)
6da7d538e5e0 changed the order of rfv and reformulated a3 with rfv
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
    28
where
6da7d538e5e0 changed the order of rfv and reformulated a3 with rfv
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
    29
  a1: "a = b \<Longrightarrow> (rVar a) \<approx> (rVar b)"
6da7d538e5e0 changed the order of rfv and reformulated a3 with rfv
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
    30
| a2: "\<lbrakk>t1 \<approx> t2; s1 \<approx> s2\<rbrakk> \<Longrightarrow> rApp t1 s1 \<approx> rApp t2 s2"
6da7d538e5e0 changed the order of rfv and reformulated a3 with rfv
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
    31
| a3: "\<lbrakk>t \<approx> ([(a,b)]\<bullet>s); a \<notin> rfv (rLam b t)\<rbrakk> \<Longrightarrow> rLam a t \<approx> rLam b s"
6da7d538e5e0 changed the order of rfv and reformulated a3 with rfv
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
    32
259
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
    33
print_theorems
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
    34
201
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
quotient lam = rlam / alpha
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
sorry
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
203
7384115df9fd added equiv-thm to the quot_info
Christian Urban <urbanc@in.tum.de>
parents: 201
diff changeset
    38
print_quotients
7384115df9fd added equiv-thm to the quot_info
Christian Urban <urbanc@in.tum.de>
parents: 201
diff changeset
    39
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    40
quotient_def (for lam)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    41
  Var :: "name \<Rightarrow> lam"
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    42
where
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    43
  "Var \<equiv> rVar"
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    44
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    45
quotient_def (for lam)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    46
  App :: "lam \<Rightarrow> lam \<Rightarrow> lam"
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    47
where
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    48
  "App \<equiv> rApp"
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    49
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    50
quotient_def (for lam)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    51
  Lam :: "name \<Rightarrow> lam \<Rightarrow> lam"
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    52
where
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    53
  "Lam \<equiv> rLam"
201
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
218
df05cd030d2f added infrastructure for defining lifted constants
Christian Urban <urbanc@in.tum.de>
parents: 203
diff changeset
    55
thm Var_def
df05cd030d2f added infrastructure for defining lifted constants
Christian Urban <urbanc@in.tum.de>
parents: 203
diff changeset
    56
thm App_def
df05cd030d2f added infrastructure for defining lifted constants
Christian Urban <urbanc@in.tum.de>
parents: 203
diff changeset
    57
thm Lam_def
df05cd030d2f added infrastructure for defining lifted constants
Christian Urban <urbanc@in.tum.de>
parents: 203
diff changeset
    58
243
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    59
quotient_def (for lam)
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    60
  fv :: "lam \<Rightarrow> name set"
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    61
where
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    62
  "fv \<equiv> rfv"
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    63
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    64
thm fv_def
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    65
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    66
(* definition of overloaded permutation function *)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    67
(* for the lifted type lam                       *)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    68
overloading
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    69
  perm_lam    \<equiv> "perm :: 'x prm \<Rightarrow> lam \<Rightarrow> lam"   (unchecked)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    70
begin
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    71
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    72
quotient_def (for lam)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    73
  perm_lam :: "'x prm \<Rightarrow> lam \<Rightarrow> lam"
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    74
where
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    75
  "perm_lam \<equiv> (perm::'x prm \<Rightarrow> rlam \<Rightarrow> rlam)"
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    76
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    77
end
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    78
238
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
    79
(*quotient_def (for lam)
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
    80
  abs_fun_lam :: "'x prm \<Rightarrow> lam \<Rightarrow> lam"
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
    81
where
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
    82
  "perm_lam \<equiv> (perm::'x prm \<Rightarrow> rlam \<Rightarrow> rlam)"*)
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
    83
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
    84
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    85
thm perm_lam_def
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    86
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    87
(* lemmas that need to lift *)
234
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
    88
lemma pi_var_com:
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    89
  fixes pi::"'x prm"
234
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
    90
  shows "(pi\<bullet>rVar a) \<approx> rVar (pi\<bullet>a)"
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    91
  sorry
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    92
234
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
    93
lemma pi_app_com:
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    94
  fixes pi::"'x prm"
234
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
    95
  shows "(pi\<bullet>rApp t1 t2) \<approx> rApp (pi\<bullet>t1) (pi\<bullet>t2)"
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    96
  sorry
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    97
234
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
    98
lemma pi_lam_com:
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    99
  fixes pi::"'x prm"
234
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   100
  shows "(pi\<bullet>rLam a t) \<approx> rLam (pi\<bullet>a) (pi\<bullet>t)"
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   101
  sorry
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   102
243
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
   103
lemma fv_var:
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
   104
  shows "fv (Var a) = {a}"
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
   105
sorry
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
   106
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
   107
lemma fv_app:
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
   108
  shows "fv (App t1 t2) = (fv t1) \<union> (fv t2)"
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
   109
sorry
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
   110
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
   111
lemma fv_lam:
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
   112
  shows "fv (Lam a t) = (fv t) - {a}"
252
e30997c88050 Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 251
diff changeset
   113
sorry
243
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
   114
201
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
lemma real_alpha:
242
47de63a883c2 The proper real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 240
diff changeset
   116
  assumes "t = [(a,b)]\<bullet>s" "a\<sharp>[b].s"
201
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
  shows "Lam a t = Lam b s"
217
9cc87d02190a First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 203
diff changeset
   118
sorry
9cc87d02190a First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 203
diff changeset
   119
9cc87d02190a First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 203
diff changeset
   120
9cc87d02190a First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 203
diff changeset
   121
9cc87d02190a First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 203
diff changeset
   122
9cc87d02190a First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 203
diff changeset
   123
9cc87d02190a First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 203
diff changeset
   124
(* Construction Site code *)
9cc87d02190a First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 203
diff changeset
   125
234
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   126
lemma perm_rsp: "(op = ===> alpha ===> alpha) op \<bullet> op \<bullet>"
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   127
  apply(auto)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   128
  (* this is propably true if some type conditions are imposed ;o) *)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   129
  sorry
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   130
234
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   131
lemma fresh_rsp: "(op = ===> alpha ===> op =) fresh fresh" 
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   132
  apply(auto)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   133
  (* this is probably only true if some type conditions are imposed *)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   134
  sorry
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   135
234
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   136
lemma rVar_rsp: "(op = ===> alpha) rVar rVar"
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   137
  apply(auto)
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   138
  apply(rule a1)
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   139
  apply(simp)
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   140
  done
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   141
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   142
lemma rApp_rsp: "(alpha ===> alpha ===> alpha) rApp rApp"
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   143
  apply(auto)
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   144
  apply(rule a2)
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   145
  apply (assumption)
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   146
  apply (assumption)
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   147
  done
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   148
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   149
lemma rLam_rsp: "(op = ===> alpha ===> alpha) rLam rLam"
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   150
  apply(auto)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   151
  apply(rule a3)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   152
  apply(rule_tac t="[(x,x)]\<bullet>y" and s="y" in subst)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   153
  apply(rule sym)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   154
  apply(rule trans)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   155
  apply(rule pt_name3)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   156
  apply(rule at_ds1[OF at_name_inst])
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   157
  apply(simp add: pt_name1)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   158
  apply(assumption)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   159
  apply(simp add: abs_fresh)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   160
  done
217
9cc87d02190a First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 203
diff changeset
   161
247
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   162
lemma rfv_rsp: "(alpha ===> op =) rfv rfv"
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   163
  sorry
217
9cc87d02190a First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 203
diff changeset
   164
9cc87d02190a First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 203
diff changeset
   165
ML {* val qty = @{typ "lam"} *}
240
6cff34032a00 Cleaning also in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 238
diff changeset
   166
ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *}
247
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   167
ML {* val defs = @{thms Var_def App_def Lam_def perm_lam_def fv_def} *}
240
6cff34032a00 Cleaning also in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 238
diff changeset
   168
ML {* val consts = lookup_quot_consts defs *}
6cff34032a00 Cleaning also in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 238
diff changeset
   169
ML {* val (trans2, reps_same, quot) = lookup_quot_thms @{context} "lam" *}
247
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   170
ML {* val rsp_thms = @{thms perm_rsp fresh_rsp rVar_rsp rApp_rsp rLam_rsp rfv_rsp} @ @{thms ho_all_prs ho_ex_prs} *}
240
6cff34032a00 Cleaning also in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 238
diff changeset
   171
ML {* fun lift_thm_lam lthy t = lift_thm lthy qty "lam" rsp_thms defs t *}
237
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   172
249
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   173
ML {* val pi_var = lift_thm_lam @{context} @{thm pi_var_com} *}
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   174
ML {* val pi_app = lift_thm_lam @{context} @{thm pi_app_com} *}
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   175
ML {* val pi_lam = lift_thm_lam @{context} @{thm pi_lam_com} *}
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   176
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   177
ML {* val fv_var = lift_thm_lam @{context} @{thm rfv_var} *}
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   178
ML {* val fv_app = lift_thm_lam @{context} @{thm rfv_app} *}
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   179
ML {* val fv_lam = lift_thm_lam @{context} @{thm rfv_lam} *}
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   180
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   181
ML {* val a1 = lift_thm_lam @{context} @{thm a1} *}
259
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   182
ML {* val a2 = lift_thm_lam @{context} @{thm a2} *}
249
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   183
ML {* val a3 = lift_thm_lam @{context} @{thm a3} *}
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   184
252
e30997c88050 Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 251
diff changeset
   185
ML {* val alpha_cases = lift_thm_lam @{context} @{thm alpha.cases} *}
e30997c88050 Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 251
diff changeset
   186
249
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   187
local_setup {*
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   188
  Quotient.note (@{binding "pi_var"}, pi_var) #> snd #>
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   189
  Quotient.note (@{binding "pi_app"}, pi_app) #> snd #>
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   190
  Quotient.note (@{binding "pi_lam"}, pi_lam) #> snd #>
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   191
  Quotient.note (@{binding "a1"}, a1) #> snd #>
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   192
  Quotient.note (@{binding "a2"}, a2) #> snd #>
252
e30997c88050 Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 251
diff changeset
   193
  Quotient.note (@{binding "a3"}, a3) #> snd #>
e30997c88050 Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 251
diff changeset
   194
  Quotient.note (@{binding "alpha_cases"}, alpha_cases) #> snd
249
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   195
*}
237
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   196
249
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   197
thm alpha.cases
252
e30997c88050 Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 251
diff changeset
   198
thm alpha_cases
249
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   199
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   200
thm rlam.inject
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   201
thm pi_var
253
e169a99c6ada Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 252
diff changeset
   202
 
249
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   203
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   204
lemma var_inject:
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   205
  shows "(Var a = Var b) = (a = b)"
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   206
sorry
247
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   207
249
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   208
lemma var_supp:
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   209
  shows "supp (Var a) = ((supp a)::name set)"
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   210
  apply(simp add: supp_def)
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   211
  apply(simp add: pi_var)
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   212
  apply(simp add: var_inject)
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   213
  done
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   214
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   215
lemma var_fresh:
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   216
  fixes a::"name"
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   217
  shows "(a\<sharp>(Var b)) = (a\<sharp>b)"
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   218
  apply(simp add: fresh_def)
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   219
  apply(simp add: var_supp)
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   220
  done
247
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   221
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   222
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   223
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   224
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   225
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   226
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   227
253
e169a99c6ada Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 252
diff changeset
   228
ML {* val t_a = atomize_thm @{thm alpha.induct} *}
e169a99c6ada Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 252
diff changeset
   229
(* prove {* build_regularize_goal t_a rty rel @{context}  *}
e169a99c6ada Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 252
diff changeset
   230
ML_prf {*  fun tac ctxt =
e169a99c6ada Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 252
diff changeset
   231
     (FIRST' [
e169a99c6ada Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 252
diff changeset
   232
      rtac rel_refl,
e169a99c6ada Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 252
diff changeset
   233
      atac,
e169a99c6ada Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 252
diff changeset
   234
      rtac @{thm universal_twice},
e169a99c6ada Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 252
diff changeset
   235
      (rtac @{thm impI} THEN' atac),
e169a99c6ada Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 252
diff changeset
   236
      rtac @{thm implication_twice},
e169a99c6ada Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 252
diff changeset
   237
      EqSubst.eqsubst_tac ctxt [0]
e169a99c6ada Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 252
diff changeset
   238
        [(@{thm equiv_res_forall} OF [rel_eqv]),
e169a99c6ada Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 252
diff changeset
   239
         (@{thm equiv_res_exists} OF [rel_eqv])],
e169a99c6ada Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 252
diff changeset
   240
      (rtac @{thm impI} THEN' (asm_full_simp_tac (Simplifier.context ctxt HOL_ss)) THEN' rtac rel_refl),
e169a99c6ada Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 252
diff changeset
   241
      (rtac @{thm RIGHT_RES_FORALL_REGULAR})
e169a99c6ada Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 252
diff changeset
   242
     ]);
e169a99c6ada Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 252
diff changeset
   243
 *}
e169a99c6ada Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 252
diff changeset
   244
  apply (tactic {* tac @{context} 1 *}) *)
251
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   245
ML {* val t_r = regularize t_a rty rel rel_eqv rel_refl @{context} *}
259
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   246
258
93ea455b29f1 Map does not fully work yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 257
diff changeset
   247
(*ML {*
257
68bd5c2a1b96 Fixed quotdata_lookup.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 253
diff changeset
   248
  val rt = build_repabs_term @{context} t_r consts rty qty
68bd5c2a1b96 Fixed quotdata_lookup.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 253
diff changeset
   249
  val rg = Logic.mk_equals ((Thm.prop_of t_r), rt);
68bd5c2a1b96 Fixed quotdata_lookup.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 253
diff changeset
   250
*}
68bd5c2a1b96 Fixed quotdata_lookup.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 253
diff changeset
   251
prove rg
68bd5c2a1b96 Fixed quotdata_lookup.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 253
diff changeset
   252
apply(atomize(full))
68bd5c2a1b96 Fixed quotdata_lookup.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 253
diff changeset
   253
ML_prf {*
68bd5c2a1b96 Fixed quotdata_lookup.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 253
diff changeset
   254
fun r_mk_comb_tac ctxt rty quot_thm reflex_thm trans_thm rsp_thms =
68bd5c2a1b96 Fixed quotdata_lookup.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 253
diff changeset
   255
  (FIRST' [
68bd5c2a1b96 Fixed quotdata_lookup.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 253
diff changeset
   256
    rtac trans_thm,
68bd5c2a1b96 Fixed quotdata_lookup.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 253
diff changeset
   257
    LAMBDA_RES_TAC ctxt,
68bd5c2a1b96 Fixed quotdata_lookup.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 253
diff changeset
   258
    res_forall_rsp_tac ctxt,
68bd5c2a1b96 Fixed quotdata_lookup.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 253
diff changeset
   259
    res_exists_rsp_tac ctxt,
68bd5c2a1b96 Fixed quotdata_lookup.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 253
diff changeset
   260
    (
68bd5c2a1b96 Fixed quotdata_lookup.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 253
diff changeset
   261
     (simp_tac ((Simplifier.context ctxt HOL_ss) addsimps rsp_thms))
68bd5c2a1b96 Fixed quotdata_lookup.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 253
diff changeset
   262
     THEN_ALL_NEW (fn _ => no_tac)
68bd5c2a1b96 Fixed quotdata_lookup.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 253
diff changeset
   263
    ),
68bd5c2a1b96 Fixed quotdata_lookup.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 253
diff changeset
   264
    (instantiate_tac @{thm REP_ABS_RSP(1)} ctxt THEN' (RANGE [quotient_tac quot_thm])),
68bd5c2a1b96 Fixed quotdata_lookup.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 253
diff changeset
   265
    rtac refl,
68bd5c2a1b96 Fixed quotdata_lookup.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 253
diff changeset
   266
    (APPLY_RSP_TAC rty ctxt THEN' (RANGE [quotient_tac quot_thm, quotient_tac quot_thm])),
68bd5c2a1b96 Fixed quotdata_lookup.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 253
diff changeset
   267
    Cong_Tac.cong_tac @{thm cong},
68bd5c2a1b96 Fixed quotdata_lookup.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 253
diff changeset
   268
    rtac @{thm ext},
68bd5c2a1b96 Fixed quotdata_lookup.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 253
diff changeset
   269
    rtac reflex_thm,
68bd5c2a1b96 Fixed quotdata_lookup.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 253
diff changeset
   270
    atac,
68bd5c2a1b96 Fixed quotdata_lookup.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 253
diff changeset
   271
    (
68bd5c2a1b96 Fixed quotdata_lookup.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 253
diff changeset
   272
     (simp_tac ((Simplifier.context ctxt HOL_ss) addsimps @{thms FUN_REL.simps}))
68bd5c2a1b96 Fixed quotdata_lookup.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 253
diff changeset
   273
     THEN_ALL_NEW (fn _ => no_tac)
68bd5c2a1b96 Fixed quotdata_lookup.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 253
diff changeset
   274
    ),
68bd5c2a1b96 Fixed quotdata_lookup.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 253
diff changeset
   275
    WEAK_LAMBDA_RES_TAC ctxt
68bd5c2a1b96 Fixed quotdata_lookup.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 253
diff changeset
   276
    ]);
68bd5c2a1b96 Fixed quotdata_lookup.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 253
diff changeset
   277
  fun r_mk_comb_tac_lam lthy = r_mk_comb_tac lthy rty quot rel_refl trans2 rsp_thms
68bd5c2a1b96 Fixed quotdata_lookup.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 253
diff changeset
   278
*}
68bd5c2a1b96 Fixed quotdata_lookup.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 253
diff changeset
   279
apply (tactic {* REPEAT_ALL_NEW (r_mk_comb_tac_lam @{context}) 1 *})
258
93ea455b29f1 Map does not fully work yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 257
diff changeset
   280
*)
257
68bd5c2a1b96 Fixed quotdata_lookup.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 253
diff changeset
   281
251
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   282
ML {* val t_t = repabs @{context} t_r consts rty qty quot rel_refl trans2 rsp_thms *}
253
e169a99c6ada Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 252
diff changeset
   283
ML {* val abs = findabs rty (prop_of (atomize_thm @{thm alpha.induct})) *}
e169a99c6ada Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 252
diff changeset
   284
ML {* val aps = findaps rty (prop_of (atomize_thm @{thm alpha.induct})) *}
265
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   285
ML {* prop_of (atomize_thm @{thm alpha.induct}) *}
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   286
ML {*
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   287
  fun findall_all rty qty tm =
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   288
    case tm of
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   289
      Const (@{const_name All}, T) $ (s as (Abs(_, _, b))) =>
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   290
        let
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   291
          val tys = findall_all rty qty s
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   292
        in if needs_lift rty T then
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   293
          (( T) :: tys)
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   294
        else tys end
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   295
    | Abs(_, T, b) =>
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   296
        findall_all rty qty (subst_bound ((Free ("x", T)), b))
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   297
    | f $ a => (findall_all rty qty f) @ (findall_all rty qty a)
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   298
    | _ => [];
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   299
  fun findall rty qty tm =
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   300
    map domain_type (
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   301
      map (old_exchange_ty rty qty)
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   302
      (distinct (op =) (findall_all rty qty tm))
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   303
    )
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   304
*}
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   305
ML {* val alls = findall rty qty (prop_of (atomize_thm @{thm alpha.induct})) *}
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   306
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   307
ML {*
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   308
fun make_simp_all_prs_thm lthy quot_thm thm typ =
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   309
  let
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   310
    val (_, [lty, rty]) = dest_Type typ;
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   311
    val thy = ProofContext.theory_of lthy;
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   312
    val (lcty, rcty) = (ctyp_of thy lty, ctyp_of thy rty)
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   313
    val inst = [NONE, SOME lcty];
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   314
    val lpi = Drule.instantiate' inst [] thm;
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   315
    val tac =
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   316
      (compose_tac (false, lpi, 1)) THEN_ALL_NEW
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   317
      (quotient_tac quot_thm);
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   318
    val gc = Drule.strip_imp_concl (cprop_of lpi);
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   319
    val t = Goal.prove_internal [] gc (fn _ => tac 1)
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   320
  in
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   321
    MetaSimplifier.rewrite_rule [@{thm eq_reflection} OF @{thms id_apply}] t
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   322
  end
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   323
*}
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   324
ML {* val simp_app_prs_thms = map (make_simp_prs_thm @{context} quot @{thm APP_PRS}) aps *}
253
e169a99c6ada Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 252
diff changeset
   325
ML {* val aps = @{typ "LamEx.rlam \<Rightarrow> bool"} :: aps; *}
265
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   326
ML {* val simp_lam_prs_thms = map (make_simp_prs_thm @{context} quot @{thm LAMBDA_PRS}) abs *}
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   327
ML {* val t_l = repeat_eqsubst_thm @{context} (simp_app_prs_thms @  simp_lam_prs_thms) t_a *}
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   328
ML {* val typ = hd (alls) *}
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   329
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   330
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   331
ML {*
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   332
    val (_, [lty, rty]) = dest_Type typ;
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   333
    val thy = @{theory};
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   334
    val (lcty, rcty) = (ctyp_of thy lty, ctyp_of thy rty)
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   335
    val inst = [NONE, SOME lcty];
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   336
    val lpi = Drule.instantiate' inst [] @{thm FORALL_PRS};
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   337
    val tac =
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   338
      (compose_tac (false, lpi, 1)) THEN_ALL_NEW
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   339
      (quotient_tac quot);
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   340
    val gc = Drule.strip_imp_concl (cprop_of lpi);
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   341
*}
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   342
prove tst: {*term_of gc*}
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   343
apply (tactic {*compose_tac (false, lpi, 1) 1 *})
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   344
apply (tactic {*quotient_tac quot 1 *})
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   345
done
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   346
thm tst
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   347
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   348
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   349
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   350
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   351
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   352
ML {* val thms = (make_simp_all_prs_thm @{context} quot @{thm FORALL_PRS} o domain_type) (hd (rev alls)) *}
5f3b364d4765 Preparing infrastructure for general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 259
diff changeset
   353
ML {* val thm =
259
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   354
  @{thm FORALL_PRS[OF FUN_QUOTIENT[OF QUOTIENT_lam FUN_QUOTIENT[OF QUOTIENT_lam IDENTITY_QUOTIENT]]]} *}
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   355
ML {* val t_a = simp_allex_prs quot [thm] t_t *}
252
e30997c88050 Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 251
diff changeset
   356
ML {* val defs_sym = add_lower_defs @{context} defs; *}
259
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   357
ML {* val defs_sym_eq = map (fn x => eq_reflection OF [x]) defs_sym *}
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   358
ML t_l
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   359
ML {* val t_d0 = MetaSimplifier.rewrite_rule defs_sym_eq t_l *}
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   360
ML {* val t_d = repeat_eqsubst_thm @{context} defs_sym t_d0 *}
252
e30997c88050 Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 251
diff changeset
   361
ML {* val t_r = MetaSimplifier.rewrite_rule [reps_same] t_d *}
253
e169a99c6ada Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 252
diff changeset
   362
ML {* val t_r1 = repeat_eqsubst_thm @{context} @{thms fun_map.simps} t_r *}
259
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   363
ML {* val t_r2 = MetaSimplifier.rewrite_rule @{thms QUOT_TYPE_I_lam.thm10} t_r1 *}
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   364
ML {* val t_r3 = MetaSimplifier.rewrite_rule @{thms eq_reflection[OF id_apply]} t_r2 *}
253
e169a99c6ada Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 252
diff changeset
   365
ML {* val alpha_induct = ObjectLogic.rulify t_r3 *}
247
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   366
253
e169a99c6ada Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 252
diff changeset
   367
local_setup {*
e169a99c6ada Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 252
diff changeset
   368
  Quotient.note (@{binding "alpha_induct"}, alpha_induct) #> snd
e169a99c6ada Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 252
diff changeset
   369
*}
247
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   370
259
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   371
thm alpha_induct
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   372
thm alpha.induct
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   373
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   374
lemma rvar_inject: "rVar a \<approx> rVar b \<Longrightarrow> (a = b)"
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   375
apply (rule alpha.cases)
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   376
apply (assumption)
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   377
apply (assumption)
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   378
apply (simp_all)
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   379
apply (cases "a = b")
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   380
apply (simp_all)
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   381
apply (cases "ba = b")
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   382
apply (simp_all)
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   383
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   384
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   385
lemma var_inject_test:
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   386
  fixes a b
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   387
  assumes a: "Var a = Var b"
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   388
  shows "(a = b)"
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   389
  using a   apply (cases "a = b")
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   390
  apply (simp_all)
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   391
  apply (rule_tac x="Var a" and xa = "Var b" in alpha_cases)
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   392
  apply (rule a)
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   393
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   394
253
e169a99c6ada Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 252
diff changeset
   395
lemma
259
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   396
  assumes a: "(x::lam) = y"
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   397
  shows "P x y"
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   398
  apply (induct rule: alpha_induct)
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   399
  apply (rule a)
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   400
thm alpha.induct
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   401
thm alpha_induct
237
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   402
fun
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   403
  option_map::"('a \<Rightarrow> 'b) \<Rightarrow> ('a noption) \<Rightarrow> ('b noption)"
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   404
where
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   405
  "option_map f (nSome x) = nSome (f x)"
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   406
| "option_map f nNone = nNone"
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   407
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   408
fun
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   409
  option_rel
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   410
where
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   411
  "option_rel r (nSome x) (nSome y) = r x y"
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   412
| "option_rel r _ _ = False"
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   413
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   414
declare [[map noption = (option_map, option_rel)]]
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   415
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   416
lemma OPT_QUOTIENT:
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   417
  assumes q: "QUOTIENT R Abs Rep"
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   418
  shows "QUOTIENT (option_rel R) (option_map Abs) (option_map Rep)"
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   419
  apply (unfold QUOTIENT_def)
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   420
  apply (auto)
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   421
  using q
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   422
  apply (unfold QUOTIENT_def)
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   423
  apply (case_tac "a :: 'b noption")
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   424
  apply (simp)
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   425
  apply (simp)
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   426
  apply (case_tac "a :: 'b noption")
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   427
  apply (simp only: option_map.simps)
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   428
  apply (subst option_rel.simps)
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   429
  (* Simp starts hanging so don't know how to continue *)
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   430
  sorry
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   431