Quot/Examples/LamEx.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Mon, 15 Feb 2010 16:51:30 +0100
changeset 1153 2ad8e66de294
parent 1139 c4001cda9da3
permissions -rw-r--r--
Fixed the definition of less and finished the missing proof.
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
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1114
diff changeset
     2
imports Nominal "../Quotient" "../Quotient_List"
201
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
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
     7
datatype rlam =
201
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
  rVar "name"
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
| rApp "rlam" "rlam"
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
| rLam "name" "rlam"
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    12
fun
243
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    13
  rfv :: "rlam \<Rightarrow> name set"
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    14
where
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    15
  rfv_var: "rfv (rVar a) = {a}"
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    16
| 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
    17
| rfv_lam: "rfv (rLam a t) = (rfv t) - {a}"
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    18
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    19
overloading
876
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 804
diff changeset
    20
  perm_rlam \<equiv> "perm :: 'x prm \<Rightarrow> rlam \<Rightarrow> rlam"   (unchecked)
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    21
begin
243
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    22
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    23
fun
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    24
  perm_rlam
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    25
where
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    26
  "perm_rlam pi (rVar a) = rVar (pi \<bullet> a)"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    27
| "perm_rlam pi (rApp t1 t2) = rApp (perm_rlam pi t1) (perm_rlam pi t2)"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    28
| "perm_rlam pi (rLam a t) = rLam (pi \<bullet> a) (perm_rlam pi t)"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    29
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    30
end
243
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    31
897
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    32
declare perm_rlam.simps[eqvt]
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    33
895
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
    34
instance rlam::pt_name
900
3bd2847cfda7 A version of hom with quantifiers.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 898
diff changeset
    35
  apply(default)
3bd2847cfda7 A version of hom with quantifiers.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 898
diff changeset
    36
  apply(induct_tac [!] x rule: rlam.induct)
3bd2847cfda7 A version of hom with quantifiers.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 898
diff changeset
    37
  apply(simp_all add: pt_name2 pt_name3)
3bd2847cfda7 A version of hom with quantifiers.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 898
diff changeset
    38
  done
895
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
    39
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
    40
instance rlam::fs_name
900
3bd2847cfda7 A version of hom with quantifiers.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 898
diff changeset
    41
  apply(default)
3bd2847cfda7 A version of hom with quantifiers.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 898
diff changeset
    42
  apply(induct_tac [!] x rule: rlam.induct)
3bd2847cfda7 A version of hom with quantifiers.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 898
diff changeset
    43
  apply(simp add: supp_def)
3bd2847cfda7 A version of hom with quantifiers.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 898
diff changeset
    44
  apply(fold supp_def)
3bd2847cfda7 A version of hom with quantifiers.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 898
diff changeset
    45
  apply(simp add: supp_atm)
3bd2847cfda7 A version of hom with quantifiers.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 898
diff changeset
    46
  apply(simp add: supp_def Collect_imp_eq Collect_neg_eq)
3bd2847cfda7 A version of hom with quantifiers.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 898
diff changeset
    47
  apply(simp add: supp_def)
3bd2847cfda7 A version of hom with quantifiers.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 898
diff changeset
    48
  apply(simp add: supp_def Collect_imp_eq Collect_neg_eq[symmetric])
3bd2847cfda7 A version of hom with quantifiers.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 898
diff changeset
    49
  apply(fold supp_def)
3bd2847cfda7 A version of hom with quantifiers.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 898
diff changeset
    50
  apply(simp add: supp_atm)
3bd2847cfda7 A version of hom with quantifiers.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 898
diff changeset
    51
  done
895
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
    52
897
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    53
declare set_diff_eqvt[eqvt]
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    54
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    55
lemma rfv_eqvt[eqvt]:
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    56
  fixes pi::"name prm"
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    57
  shows "(pi\<bullet>rfv t) = rfv (pi\<bullet>t)"
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    58
apply(induct t)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    59
apply(simp_all)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    60
apply(simp add: perm_set_eq)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    61
apply(simp add: union_eqvt)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    62
apply(simp add: set_diff_eqvt)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    63
apply(simp add: perm_set_eq)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    64
done
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    65
271
1b57f99737fe Alpha.induct now lifts automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 270
diff changeset
    66
inductive
897
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    67
    alpha :: "rlam \<Rightarrow> rlam \<Rightarrow> bool" ("_ \<approx> _" [100, 100] 100)
246
6da7d538e5e0 changed the order of rfv and reformulated a3 with rfv
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
    68
where
6da7d538e5e0 changed the order of rfv and reformulated a3 with rfv
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
    69
  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
    70
| a2: "\<lbrakk>t1 \<approx> t2; s1 \<approx> s2\<rbrakk> \<Longrightarrow> rApp t1 s1 \<approx> rApp t2 s2"
915
16082d0b8ac1 Trying to define hom for the lifted type directly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 903
diff changeset
    71
| a3: "\<exists>pi::name prm. (rfv t - {a} = rfv s - {b} \<and> (rfv t - {a})\<sharp>* pi \<and> (pi \<bullet> t) \<approx> s \<and> (pi \<bullet> a) = b)
897
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    72
       \<Longrightarrow> rLam a t \<approx> rLam b s"
246
6da7d538e5e0 changed the order of rfv and reformulated a3 with rfv
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
    73
918
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 916
diff changeset
    74
897
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    75
(* should be automatic with new version of eqvt-machinery *)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    76
lemma alpha_eqvt:
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    77
  fixes pi::"name prm"
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    78
  shows "t \<approx> s \<Longrightarrow> (pi \<bullet> t) \<approx> (pi \<bullet> s)"
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    79
apply(induct rule: alpha.induct)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    80
apply(simp add: a1)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    81
apply(simp add: a2)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    82
apply(simp)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    83
apply(rule a3)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    84
apply(erule conjE)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    85
apply(erule exE)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    86
apply(erule conjE)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    87
apply(rule_tac x="pi \<bullet> pia" in exI)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    88
apply(rule conjI)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    89
apply(rule_tac pi1="rev pi" in perm_bij[THEN iffD1])
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    90
apply(perm_simp add: eqvts)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    91
apply(rule conjI)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    92
apply(rule_tac pi1="rev pi" in pt_fresh_star_bij(1)[OF pt_name_inst at_name_inst, THEN iffD1])
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    93
apply(perm_simp add: eqvts)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    94
apply(rule conjI)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    95
apply(subst perm_compose[symmetric])
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    96
apply(simp)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    97
apply(subst perm_compose[symmetric])
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    98
apply(simp)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
    99
done
259
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
   100
271
1b57f99737fe Alpha.induct now lifts automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 270
diff changeset
   101
lemma alpha_refl:
286
Christian Urban <urbanc@in.tum.de>
parents: 285
diff changeset
   102
  shows "t \<approx> t"
897
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   103
apply(induct t rule: rlam.induct)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   104
apply(simp add: a1)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   105
apply(simp add: a2)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   106
apply(rule a3)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   107
apply(rule_tac x="[]" in exI)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   108
apply(simp_all add: fresh_star_def fresh_list_nil)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   109
done
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   110
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   111
lemma alpha_sym:
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   112
  shows "t \<approx> s \<Longrightarrow> s \<approx> t"
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   113
apply(induct rule: alpha.induct)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   114
apply(simp add: a1)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   115
apply(simp add: a2)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   116
apply(rule a3)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   117
apply(erule exE)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   118
apply(rule_tac x="rev pi" in exI)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   119
apply(simp)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   120
apply(simp add: fresh_star_def fresh_list_rev)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   121
apply(rule conjI)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   122
apply(erule conjE)+
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   123
apply(rotate_tac 3)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   124
apply(drule_tac pi="rev pi" in alpha_eqvt)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   125
apply(perm_simp)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   126
apply(rule pt_bij2[OF pt_name_inst at_name_inst])
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   127
apply(simp)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   128
done
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   129
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   130
lemma alpha_trans:
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   131
  shows "t1 \<approx> t2 \<Longrightarrow> t2 \<approx> t3 \<Longrightarrow> t1 \<approx> t3"
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   132
apply(induct arbitrary: t3 rule: alpha.induct)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   133
apply(erule alpha.cases)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   134
apply(simp_all)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   135
apply(simp add: a1)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   136
apply(rotate_tac 4)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   137
apply(erule alpha.cases)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   138
apply(simp_all)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   139
apply(simp add: a2)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   140
apply(rotate_tac 1)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   141
apply(erule alpha.cases)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   142
apply(simp_all)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   143
apply(erule conjE)+
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   144
apply(erule exE)+
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   145
apply(erule conjE)+
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   146
apply(rule a3)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   147
apply(rule_tac x="pia @ pi" in exI)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   148
apply(simp add: fresh_star_def fresh_list_append)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   149
apply(simp add: pt_name2)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   150
apply(drule_tac x="rev pia \<bullet> sa" in spec)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   151
apply(drule mp)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   152
apply(rotate_tac 8)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   153
apply(drule_tac pi="rev pia" in alpha_eqvt)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   154
apply(perm_simp)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   155
apply(rotate_tac 11)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   156
apply(drule_tac pi="pia" in alpha_eqvt)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   157
apply(perm_simp)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   158
done
286
Christian Urban <urbanc@in.tum.de>
parents: 285
diff changeset
   159
534
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 515
diff changeset
   160
lemma alpha_equivp:
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 515
diff changeset
   161
  shows "equivp alpha"
897
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   162
apply(rule equivpI)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   163
unfolding reflp_def symp_def transp_def
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   164
apply(auto intro: alpha_refl alpha_sym alpha_trans)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   165
done
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   166
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   167
lemma alpha_rfv:
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   168
  shows "t \<approx> s \<Longrightarrow> rfv t = rfv s"
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   169
apply(induct rule: alpha.induct)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   170
apply(simp)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   171
apply(simp)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   172
apply(simp)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   173
done
271
1b57f99737fe Alpha.induct now lifts automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 270
diff changeset
   174
766
df053507edba renamed "quotient" command to "quotient_type"; needs new keyword file to be installed
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   175
quotient_type lam = rlam / alpha
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   176
  by (rule alpha_equivp)
201
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
203
7384115df9fd added equiv-thm to the quot_info
Christian Urban <urbanc@in.tum.de>
parents: 201
diff changeset
   178
767
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 766
diff changeset
   179
quotient_definition
876
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 804
diff changeset
   180
  "Var :: name \<Rightarrow> lam"
1139
c4001cda9da3 renamed 'as' to 'is' everywhere.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1128
diff changeset
   181
is
663
0dd10a900cae Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 636
diff changeset
   182
  "rVar"
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   183
767
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 766
diff changeset
   184
quotient_definition
705
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 663
diff changeset
   185
   "App :: lam \<Rightarrow> lam \<Rightarrow> lam"
1139
c4001cda9da3 renamed 'as' to 'is' everywhere.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1128
diff changeset
   186
is
663
0dd10a900cae Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 636
diff changeset
   187
  "rApp"
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   188
767
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 766
diff changeset
   189
quotient_definition
876
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 804
diff changeset
   190
  "Lam :: name \<Rightarrow> lam \<Rightarrow> lam"
1139
c4001cda9da3 renamed 'as' to 'is' everywhere.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1128
diff changeset
   191
is
663
0dd10a900cae Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 636
diff changeset
   192
  "rLam"
201
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   193
767
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 766
diff changeset
   194
quotient_definition
876
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 804
diff changeset
   195
  "fv :: lam \<Rightarrow> name set"
1139
c4001cda9da3 renamed 'as' to 'is' everywhere.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1128
diff changeset
   196
is
663
0dd10a900cae Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 636
diff changeset
   197
  "rfv"
243
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
   198
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   199
(* 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
   200
(* 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
   201
overloading
268
4d58c02289ca simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
Christian Urban <urbanc@in.tum.de>
parents: 267
diff changeset
   202
  perm_lam \<equiv> "perm :: 'x prm \<Rightarrow> lam \<Rightarrow> lam"   (unchecked)
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   203
begin
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   204
767
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 766
diff changeset
   205
quotient_definition
896
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   206
  "perm_lam :: 'x prm \<Rightarrow> lam \<Rightarrow> lam"
1139
c4001cda9da3 renamed 'as' to 'is' everywhere.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1128
diff changeset
   207
is
663
0dd10a900cae Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 636
diff changeset
   208
  "perm::'x prm \<Rightarrow> rlam \<Rightarrow> rlam"
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   209
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   210
end
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   211
636
520a4084d064 changed names of attributes
Christian Urban <urbanc@in.tum.de>
parents: 610
diff changeset
   212
lemma perm_rsp[quot_respect]:
286
Christian Urban <urbanc@in.tum.de>
parents: 285
diff changeset
   213
  "(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
   214
  apply(auto)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   215
  (* 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
   216
  sorry
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   217
451
586e3dc4afdb Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 445
diff changeset
   218
lemma fresh_rsp:
586e3dc4afdb Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 445
diff changeset
   219
  "(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
   220
  apply(auto)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   221
  (* 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
   222
  sorry
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   223
636
520a4084d064 changed names of attributes
Christian Urban <urbanc@in.tum.de>
parents: 610
diff changeset
   224
lemma rVar_rsp[quot_respect]:
451
586e3dc4afdb Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 445
diff changeset
   225
  "(op = ===> alpha) rVar rVar"
876
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 804
diff changeset
   226
  by (auto intro: a1)
234
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   227
636
520a4084d064 changed names of attributes
Christian Urban <urbanc@in.tum.de>
parents: 610
diff changeset
   228
lemma rApp_rsp[quot_respect]: "(alpha ===> alpha ===> alpha) rApp rApp"
876
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 804
diff changeset
   229
  by (auto intro: a2)
234
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   230
636
520a4084d064 changed names of attributes
Christian Urban <urbanc@in.tum.de>
parents: 610
diff changeset
   231
lemma rLam_rsp[quot_respect]: "(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
   232
  apply(auto)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   233
  apply(rule a3)
897
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   234
  apply(rule_tac x="[]" in exI)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   235
  unfolding fresh_star_def
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   236
  apply(simp add: fresh_list_nil)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   237
  apply(simp add: alpha_rfv)
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   238
  done
217
9cc87d02190a First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 203
diff changeset
   239
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   240
lemma rfv_rsp[quot_respect]: 
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   241
  "(alpha ===> op =) rfv rfv"
897
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   242
apply(simp add: alpha_rfv)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   243
done
217
9cc87d02190a First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 203
diff changeset
   244
897
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   245
section {* lifted theorems *}
370
09e28d4c19aa Lambda & SOLVED' for new quotient_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   246
896
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   247
lemma lam_induct:
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   248
  "\<lbrakk>\<And>name. P (Var name);
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   249
    \<And>lam1 lam2. \<lbrakk>P lam1; P lam2\<rbrakk> \<Longrightarrow> P (App lam1 lam2);
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   250
    \<And>name lam. P lam \<Longrightarrow> P (Lam name lam)\<rbrakk> 
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   251
    \<Longrightarrow> P lam"
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   252
  by (lifting rlam.induct)
249
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   253
1114
Christian Urban <urbanc@in.tum.de>
parents: 918
diff changeset
   254
ML {* show_all_types := true *}
Christian Urban <urbanc@in.tum.de>
parents: 918
diff changeset
   255
896
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   256
lemma perm_lam [simp]:
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   257
  fixes pi::"'a prm"
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   258
  shows "pi \<bullet> Var a = Var (pi \<bullet> a)"
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   259
  and   "pi \<bullet> App t1 t2 = App (pi \<bullet> t1) (pi \<bullet> t2)"
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   260
  and   "pi \<bullet> Lam a t = Lam (pi \<bullet> a) (pi \<bullet> t)"
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   261
apply(lifting perm_rlam.simps)
1114
Christian Urban <urbanc@in.tum.de>
parents: 918
diff changeset
   262
ML_prf {*
Christian Urban <urbanc@in.tum.de>
parents: 918
diff changeset
   263
  List.last (map (symmetric o #def) (Quotient_Info.qconsts_dest @{context}));
Christian Urban <urbanc@in.tum.de>
parents: 918
diff changeset
   264
  List.last (map (Thm.varifyT o symmetric o #def) (Quotient_Info.qconsts_dest @{context}))
Christian Urban <urbanc@in.tum.de>
parents: 918
diff changeset
   265
*}
896
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   266
done
376
e99c0334d8bf lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 370
diff changeset
   267
896
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   268
instance lam::pt_name
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   269
apply(default)
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   270
apply(induct_tac [!] x rule: lam_induct)
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   271
apply(simp_all add: pt_name2 pt_name3)
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   272
done
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   273
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   274
lemma fv_lam [simp]: 
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   275
  shows "fv (Var a) = {a}"
896
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   276
  and   "fv (App t1 t2) = fv t1 \<union> fv t2"
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   277
  and   "fv (Lam a t) = fv t - {a}"
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   278
apply(lifting rfv_var rfv_app rfv_lam)
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   279
done
249
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   280
376
e99c0334d8bf lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 370
diff changeset
   281
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   282
lemma a1: 
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   283
  "a = b \<Longrightarrow> Var a = Var b"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   284
  by  (lifting a1)
376
e99c0334d8bf lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 370
diff changeset
   285
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   286
lemma a2: 
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   287
  "\<lbrakk>x = xa; xb = xc\<rbrakk> \<Longrightarrow> App x xb = App xa xc"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   288
  by  (lifting a2)
376
e99c0334d8bf lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 370
diff changeset
   289
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   290
lemma a3: 
897
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   291
  "\<lbrakk>\<exists>pi::name prm. (fv t - {a} = fv s - {b} \<and> (fv t - {a})\<sharp>* pi \<and> (pi \<bullet> t) = s \<and> (pi \<bullet> a) = b)\<rbrakk> 
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   292
   \<Longrightarrow> Lam a t = Lam b s"
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   293
  by  (lifting a3)
286
Christian Urban <urbanc@in.tum.de>
parents: 285
diff changeset
   294
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   295
lemma alpha_cases: 
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   296
  "\<lbrakk>a1 = a2; \<And>a b. \<lbrakk>a1 = Var a; a2 = Var b; a = b\<rbrakk> \<Longrightarrow> P;
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   297
    \<And>x xa xb xc. \<lbrakk>a1 = App x xb; a2 = App xa xc; x = xa; xb = xc\<rbrakk> \<Longrightarrow> P;
897
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   298
    \<And>t a s b. \<lbrakk>a1 = Lam a t; a2 = Lam b s; 
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   299
         \<exists>pi::name prm. fv t - {a} = fv s - {b} \<and> (fv t - {a}) \<sharp>* pi \<and> (pi \<bullet> t) = s \<and> pi \<bullet> a = b\<rbrakk> \<Longrightarrow> P\<rbrakk>
370
09e28d4c19aa Lambda & SOLVED' for new quotient_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   300
    \<Longrightarrow> P"
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   301
  by (lifting alpha.cases)
378
86fba2c4eeef All examples work again.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 376
diff changeset
   302
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   303
lemma alpha_induct: 
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   304
  "\<lbrakk>qx = qxa; \<And>a b. a = b \<Longrightarrow> qxb (Var a) (Var b);
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   305
    \<And>x xa xb xc. \<lbrakk>x = xa; qxb x xa; xb = xc; qxb xb xc\<rbrakk> \<Longrightarrow> qxb (App x xb) (App xa xc);
897
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   306
     \<And>t a s b.
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   307
        \<lbrakk>\<exists>pi::name prm. fv t - {a} = fv s - {b} \<and>
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   308
         (fv t - {a}) \<sharp>* pi \<and> ((pi \<bullet> t) = s \<and> qxb (pi \<bullet> t) s) \<and> pi \<bullet> a = b\<rbrakk> \<Longrightarrow> qxb (Lam a t) (Lam b s)\<rbrakk>
370
09e28d4c19aa Lambda & SOLVED' for new quotient_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   309
    \<Longrightarrow> qxb qx qxa"
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   310
  by (lifting alpha.induct)
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   311
897
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   312
lemma lam_inject [simp]: 
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   313
  shows "(Var a = Var b) = (a = b)"
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   314
  and   "(App t1 t2 = App s1 s2) = (t1 = s1 \<and> t2 = s2)"
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   315
apply(lifting rlam.inject(1) rlam.inject(2))
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   316
apply(auto)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   317
apply(drule alpha.cases)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   318
apply(simp_all)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   319
apply(simp add: alpha.a1)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   320
apply(drule alpha.cases)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   321
apply(simp_all)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   322
apply(drule alpha.cases)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   323
apply(simp_all)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   324
apply(rule alpha.a2)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   325
apply(simp_all)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   326
done
883
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   327
916
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   328
lemma rlam_distinct:
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   329
  shows "\<not>(rVar nam \<approx> rApp rlam1' rlam2')"
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   330
  and   "\<not>(rApp rlam1' rlam2' \<approx> rVar nam)"
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   331
  and   "\<not>(rVar nam \<approx> rLam nam' rlam')"
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   332
  and   "\<not>(rLam nam' rlam' \<approx> rVar nam)"
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   333
  and   "\<not>(rApp rlam1 rlam2 \<approx> rLam nam' rlam')"
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   334
  and   "\<not>(rLam nam' rlam' \<approx> rApp rlam1 rlam2)"
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   335
apply auto
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   336
apply(erule alpha.cases)
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   337
apply simp_all
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   338
apply(erule alpha.cases)
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   339
apply simp_all
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   340
apply(erule alpha.cases)
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   341
apply simp_all
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   342
apply(erule alpha.cases)
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   343
apply simp_all
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   344
apply(erule alpha.cases)
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   345
apply simp_all
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   346
apply(erule alpha.cases)
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   347
apply simp_all
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   348
done
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   349
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   350
lemma lam_distinct[simp]:
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   351
  shows "Var nam \<noteq> App lam1' lam2'"
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   352
  and   "App lam1' lam2' \<noteq> Var nam"
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   353
  and   "Var nam \<noteq> Lam nam' lam'"
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   354
  and   "Lam nam' lam' \<noteq> Var nam"
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   355
  and   "App lam1 lam2 \<noteq> Lam nam' lam'"
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   356
  and   "Lam nam' lam' \<noteq> App lam1 lam2"
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   357
apply(lifting rlam_distinct(1) rlam_distinct(2) rlam_distinct(3) rlam_distinct(4) rlam_distinct(5) rlam_distinct(6))
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   358
done
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   359
883
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   360
lemma var_supp1:
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   361
  shows "(supp (Var a)) = ((supp a)::name set)"
916
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   362
  by (simp add: supp_def)
883
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   363
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   364
lemma var_supp:
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   365
  shows "(supp (Var a)) = {a::name}"
916
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   366
  using var_supp1 by (simp add: supp_atm)
883
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   367
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   368
lemma app_supp:
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   369
  shows "supp (App t1 t2) = (supp t1) \<union> ((supp t2)::name set)"
897
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   370
apply(simp only: perm_lam supp_def lam_inject)
883
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   371
apply(simp add: Collect_imp_eq Collect_neg_eq)
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   372
done
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   373
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   374
lemma lam_supp:
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   375
  shows "supp (Lam x t) = ((supp ([x].t))::name set)"
896
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   376
apply(simp add: supp_def)
897
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   377
apply(simp add: abs_perm)
883
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   378
sorry
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   379
878
c3662f845129 setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents: 877
diff changeset
   380
c3662f845129 setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents: 877
diff changeset
   381
instance lam::fs_name
c3662f845129 setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents: 877
diff changeset
   382
apply(default)
883
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   383
apply(induct_tac x rule: lam_induct)
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   384
apply(simp add: var_supp)
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   385
apply(simp add: app_supp)
897
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   386
apply(simp add: lam_supp abs_supp)
464619898890 used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Christian Urban <urbanc@in.tum.de>
parents: 896
diff changeset
   387
done
877
09a64cb04851 exported absrep_const for nitpick.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 876
diff changeset
   388
881
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   389
lemma fresh_lam:
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   390
  "(a \<sharp> Lam b t) \<longleftrightarrow> (a = b) \<or> (a \<noteq> b \<and> a \<sharp> t)"
883
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   391
apply(simp add: fresh_def)
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   392
apply(simp add: lam_supp abs_supp)
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   393
apply(auto)
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   394
done
881
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   395
877
09a64cb04851 exported absrep_const for nitpick.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 876
diff changeset
   396
lemma lam_induct_strong:
878
c3662f845129 setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents: 877
diff changeset
   397
  fixes a::"'a::fs_name"
879
f2a1ebba9bdc First subgoal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 878
diff changeset
   398
  assumes a1: "\<And>name b. P b (Var name)"
f2a1ebba9bdc First subgoal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 878
diff changeset
   399
  and     a2: "\<And>lam1 lam2 b. \<lbrakk>\<And>c. P c lam1; \<And>c. P c lam2\<rbrakk> \<Longrightarrow> P b (App lam1 lam2)"
881
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   400
  and     a3: "\<And>name lam b. \<lbrakk>\<And>c. P c lam; name \<sharp> b\<rbrakk> \<Longrightarrow> P b (Lam name lam)"
879
f2a1ebba9bdc First subgoal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 878
diff changeset
   401
  shows "P a lam"
878
c3662f845129 setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents: 877
diff changeset
   402
proof -
880
cd3f1409780a right generalisation
Christian Urban <urbanc@in.tum.de>
parents: 879
diff changeset
   403
  have "\<And>(pi::name prm) a. P a (pi \<bullet> lam)" 
878
c3662f845129 setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents: 877
diff changeset
   404
  proof (induct lam rule: lam_induct)
c3662f845129 setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents: 877
diff changeset
   405
    case (1 name pi)
879
f2a1ebba9bdc First subgoal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 878
diff changeset
   406
    show "P a (pi \<bullet> Var name)"
896
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   407
      apply (simp)
879
f2a1ebba9bdc First subgoal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 878
diff changeset
   408
      apply (rule a1)
f2a1ebba9bdc First subgoal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 878
diff changeset
   409
      done
878
c3662f845129 setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents: 877
diff changeset
   410
  next
c3662f845129 setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents: 877
diff changeset
   411
    case (2 lam1 lam2 pi)
880
cd3f1409780a right generalisation
Christian Urban <urbanc@in.tum.de>
parents: 879
diff changeset
   412
    have b1: "\<And>(pi::name prm) a. P a (pi \<bullet> lam1)" by fact
cd3f1409780a right generalisation
Christian Urban <urbanc@in.tum.de>
parents: 879
diff changeset
   413
    have b2: "\<And>(pi::name prm) a. P a (pi \<bullet> lam2)" by fact
879
f2a1ebba9bdc First subgoal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 878
diff changeset
   414
    show "P a (pi \<bullet> App lam1 lam2)"
896
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   415
      apply (simp)
879
f2a1ebba9bdc First subgoal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 878
diff changeset
   416
      apply (rule a2)
880
cd3f1409780a right generalisation
Christian Urban <urbanc@in.tum.de>
parents: 879
diff changeset
   417
      apply (rule b1)
cd3f1409780a right generalisation
Christian Urban <urbanc@in.tum.de>
parents: 879
diff changeset
   418
      apply (rule b2)
cd3f1409780a right generalisation
Christian Urban <urbanc@in.tum.de>
parents: 879
diff changeset
   419
      done
878
c3662f845129 setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents: 877
diff changeset
   420
  next
881
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   421
    case (3 name lam pi a)
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   422
    have b: "\<And>(pi::name prm) a. P a (pi \<bullet> lam)" by fact
882
6a8858ba01f6 removed one sorry
Christian Urban <urbanc@in.tum.de>
parents: 881
diff changeset
   423
    obtain c::name where fr: "c\<sharp>(a, pi\<bullet>name, pi\<bullet>lam)"
6a8858ba01f6 removed one sorry
Christian Urban <urbanc@in.tum.de>
parents: 881
diff changeset
   424
      apply(rule exists_fresh[of "(a, pi\<bullet>name, pi\<bullet>lam)"])
6a8858ba01f6 removed one sorry
Christian Urban <urbanc@in.tum.de>
parents: 881
diff changeset
   425
      apply(simp_all add: fs_name1)
6a8858ba01f6 removed one sorry
Christian Urban <urbanc@in.tum.de>
parents: 881
diff changeset
   426
      done
881
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   427
    from b fr have p: "P a (Lam c (([(c, pi\<bullet>name)]@pi)\<bullet>lam))" 
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   428
      apply -
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   429
      apply(rule a3)
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   430
      apply(blast)
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   431
      apply(simp)
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   432
      done
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   433
    have eq: "[(c, pi\<bullet>name)] \<bullet> Lam (pi \<bullet> name) (pi \<bullet> lam) = Lam (pi \<bullet> name) (pi \<bullet> lam)"
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   434
      apply(rule perm_fresh_fresh)
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   435
      using fr
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   436
      apply(simp add: fresh_lam)
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   437
      apply(simp add: fresh_lam)
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   438
      done
879
f2a1ebba9bdc First subgoal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 878
diff changeset
   439
    show "P a (pi \<bullet> Lam name lam)" 
896
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   440
      apply (simp)
881
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   441
      apply(subst eq[symmetric])
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   442
      using p
896
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   443
      apply(simp only: perm_lam pt_name2 swap_simps)
881
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   444
      done
878
c3662f845129 setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents: 877
diff changeset
   445
  qed
c3662f845129 setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents: 877
diff changeset
   446
  then have "P a (([]::name prm) \<bullet> lam)" by blast
c3662f845129 setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents: 877
diff changeset
   447
  then show "P a lam" by simp 
c3662f845129 setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents: 877
diff changeset
   448
qed
877
09a64cb04851 exported absrep_const for nitpick.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 876
diff changeset
   449
09a64cb04851 exported absrep_const for nitpick.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 876
diff changeset
   450
249
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   451
lemma var_fresh:
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   452
  fixes a::"name"
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   453
  shows "(a \<sharp> (Var b)) = (a \<sharp> b)"
249
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   454
  apply(simp add: fresh_def)
884
e49c6b6f37f4 tuned quotient_def.ML and cleaned somewhat LamEx.thy
Christian Urban <urbanc@in.tum.de>
parents: 883
diff changeset
   455
  apply(simp add: var_supp1)
249
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   456
  done
247
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   457
891
7bac7dffadeb hom lifted to hom', so it is true. Infrastructure for partially regularized quantifiers. Nicer errors for regularize.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 890
diff changeset
   458
(* lemma hom_reg: *)
887
d2660637e764 Incorrect version of the homomorphism lemma
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 884
diff changeset
   459
895
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   460
lemma rlam_rec_eqvt:
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   461
  fixes pi::"name prm"
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   462
  and   f1::"name \<Rightarrow> ('a::pt_name)"
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   463
  shows "(pi\<bullet>rlam_rec f1 f2 f3 t) = rlam_rec (pi\<bullet>f1) (pi\<bullet>f2) (pi\<bullet>f3) (pi\<bullet>t)"
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   464
apply(induct t)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   465
apply(simp_all)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   466
apply(simp add: perm_fun_def)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   467
apply(perm_simp)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   468
apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   469
back
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   470
apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   471
apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   472
apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   473
apply(simp)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   474
apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   475
back
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   476
apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   477
apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   478
apply(simp)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   479
done
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   480
 
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   481
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   482
lemma rlam_rec_respects:
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   483
  assumes f1: "f_var \<in> Respects (op= ===> op=)"
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   484
  and     f2: "f_app \<in> Respects (alpha ===> alpha ===> op= ===> op= ===> op=)"
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   485
  and     f3: "f_lam \<in> Respects (op= ===> alpha ===> op= ===> op=)"
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   486
  shows "rlam_rec f_var f_app f_lam \<in> Respects (alpha ===> op =)"
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   487
apply(simp add: mem_def)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   488
apply(simp add: Respects_def)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   489
apply(rule allI)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   490
apply(rule allI)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   491
apply(rule impI)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   492
apply(erule alpha.induct)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   493
apply(simp)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   494
apply(simp)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   495
using f2
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   496
apply(simp add: mem_def)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   497
apply(simp add: Respects_def)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   498
using f3[simplified mem_def Respects_def]
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   499
apply(simp)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   500
apply(case_tac "a=b")
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   501
apply(clarify)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   502
apply(simp)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   503
(* probably true *)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   504
sorry
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   505
901
28e084a66c7f term1_hom as a function
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 900
diff changeset
   506
function
903
f7cafd3c86b0 Statement of term1_hom_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 902
diff changeset
   507
  term1_hom :: "(name \<Rightarrow> 'a) \<Rightarrow>
f7cafd3c86b0 Statement of term1_hom_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 902
diff changeset
   508
                (rlam \<Rightarrow> rlam \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow>
902
82cdc3755c2c proved that the function is a function
Christian Urban <urbanc@in.tum.de>
parents: 901
diff changeset
   509
                ((name \<Rightarrow> rlam) \<Rightarrow> (name \<Rightarrow> 'a) \<Rightarrow> 'a) \<Rightarrow> rlam \<Rightarrow> 'a"
901
28e084a66c7f term1_hom as a function
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 900
diff changeset
   510
where
28e084a66c7f term1_hom as a function
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 900
diff changeset
   511
  "term1_hom var app abs' (rVar x) = (var x)"
28e084a66c7f term1_hom as a function
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 900
diff changeset
   512
| "term1_hom var app abs' (rApp t u) =
28e084a66c7f term1_hom as a function
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 900
diff changeset
   513
     app t u (term1_hom var app abs' t) (term1_hom var app abs' u)"
28e084a66c7f term1_hom as a function
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 900
diff changeset
   514
| "term1_hom var app abs' (rLam x u) =
28e084a66c7f term1_hom as a function
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 900
diff changeset
   515
     abs' (\<lambda>y. [(x, y)] \<bullet> u) (\<lambda>y. term1_hom var app abs' ([(x, y)] \<bullet> u))"
902
82cdc3755c2c proved that the function is a function
Christian Urban <urbanc@in.tum.de>
parents: 901
diff changeset
   516
apply(pat_completeness)
82cdc3755c2c proved that the function is a function
Christian Urban <urbanc@in.tum.de>
parents: 901
diff changeset
   517
apply(auto)
82cdc3755c2c proved that the function is a function
Christian Urban <urbanc@in.tum.de>
parents: 901
diff changeset
   518
done
901
28e084a66c7f term1_hom as a function
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 900
diff changeset
   519
902
82cdc3755c2c proved that the function is a function
Christian Urban <urbanc@in.tum.de>
parents: 901
diff changeset
   520
lemma pi_size:
82cdc3755c2c proved that the function is a function
Christian Urban <urbanc@in.tum.de>
parents: 901
diff changeset
   521
  fixes pi::"name prm"
82cdc3755c2c proved that the function is a function
Christian Urban <urbanc@in.tum.de>
parents: 901
diff changeset
   522
  and   t::"rlam"
82cdc3755c2c proved that the function is a function
Christian Urban <urbanc@in.tum.de>
parents: 901
diff changeset
   523
  shows "size (pi \<bullet> t) = size t"
82cdc3755c2c proved that the function is a function
Christian Urban <urbanc@in.tum.de>
parents: 901
diff changeset
   524
apply(induct t)
82cdc3755c2c proved that the function is a function
Christian Urban <urbanc@in.tum.de>
parents: 901
diff changeset
   525
apply(auto)
82cdc3755c2c proved that the function is a function
Christian Urban <urbanc@in.tum.de>
parents: 901
diff changeset
   526
done
82cdc3755c2c proved that the function is a function
Christian Urban <urbanc@in.tum.de>
parents: 901
diff changeset
   527
82cdc3755c2c proved that the function is a function
Christian Urban <urbanc@in.tum.de>
parents: 901
diff changeset
   528
termination term1_hom
82cdc3755c2c proved that the function is a function
Christian Urban <urbanc@in.tum.de>
parents: 901
diff changeset
   529
  apply(relation "measure (\<lambda>(f1, f2, f3, t). size t)")
82cdc3755c2c proved that the function is a function
Christian Urban <urbanc@in.tum.de>
parents: 901
diff changeset
   530
apply(auto simp add: pi_size)
82cdc3755c2c proved that the function is a function
Christian Urban <urbanc@in.tum.de>
parents: 901
diff changeset
   531
done
901
28e084a66c7f term1_hom as a function
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 900
diff changeset
   532
916
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   533
lemma lam_exhaust:
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   534
  "\<lbrakk>\<And>name. y = Var name \<Longrightarrow> P; \<And>rlam1 rlam2. y = App rlam1 rlam2 \<Longrightarrow> P; \<And>name rlam. y = Lam name rlam \<Longrightarrow> P\<rbrakk>
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   535
    \<Longrightarrow> P"
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   536
apply(lifting rlam.exhaust)
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   537
done
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   538
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   539
(* THIS IS NOT TRUE, but it lets prove the existence of the hom function *)
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   540
lemma lam_inject':
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   541
  "(Lam a x = Lam b y) = ((\<lambda>c. [(a, c)] \<bullet> x) = (\<lambda>c. [(b, c)] \<bullet> y))"
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   542
sorry
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   543
915
16082d0b8ac1 Trying to define hom for the lifted type directly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 903
diff changeset
   544
function
16082d0b8ac1 Trying to define hom for the lifted type directly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 903
diff changeset
   545
  hom :: "(name \<Rightarrow> 'a) \<Rightarrow>
16082d0b8ac1 Trying to define hom for the lifted type directly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 903
diff changeset
   546
                (lam \<Rightarrow> lam \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow>
16082d0b8ac1 Trying to define hom for the lifted type directly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 903
diff changeset
   547
                ((name \<Rightarrow> lam) \<Rightarrow> (name \<Rightarrow> 'a) \<Rightarrow> 'a) \<Rightarrow> lam \<Rightarrow> 'a"
16082d0b8ac1 Trying to define hom for the lifted type directly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 903
diff changeset
   548
where
16082d0b8ac1 Trying to define hom for the lifted type directly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 903
diff changeset
   549
  "hom f_var f_app f_lam (Var x) = f_var x"
16082d0b8ac1 Trying to define hom for the lifted type directly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 903
diff changeset
   550
| "hom f_var f_app f_lam (App l r) = f_app l r (hom f_var f_app f_lam l) (hom f_var f_app f_lam r)"
16082d0b8ac1 Trying to define hom for the lifted type directly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 903
diff changeset
   551
| "hom f_var f_app f_lam (Lam a x) = f_lam (\<lambda>b. ([(a,b)] \<bullet> x)) (\<lambda>b. hom f_var f_app f_lam ([(a,b)] \<bullet> x))"
916
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   552
defer
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   553
apply(simp_all add: lam_inject') (* inject, distinct *)
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   554
apply(tactic {* Cong_Tac.cong_tac @{thm cong} 1 *})
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   555
apply(rule refl)
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   556
apply(rule ext)
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   557
apply(tactic {* Cong_Tac.cong_tac @{thm cong} 1 *})
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   558
apply simp_all
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   559
apply(erule conjE)+
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   560
apply(rule_tac x="b" in cong)
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   561
apply simp_all
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   562
apply auto
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   563
apply(rule_tac y="b" in lam_exhaust)
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   564
apply simp_all
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   565
apply auto
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   566
apply meson
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   567
apply(simp_all add: lam_inject')
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   568
apply metis
915
16082d0b8ac1 Trying to define hom for the lifted type directly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 903
diff changeset
   569
done
916
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   570
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   571
termination hom
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   572
  apply -
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   573
(*
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   574
ML_prf {* Size.size_thms @{theory} "LamEx.lam" *}
915
16082d0b8ac1 Trying to define hom for the lifted type directly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 903
diff changeset
   575
*)
916
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   576
sorry
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   577
a7bf638e9af3 More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 915
diff changeset
   578
thm hom.simps
915
16082d0b8ac1 Trying to define hom for the lifted type directly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 903
diff changeset
   579
903
f7cafd3c86b0 Statement of term1_hom_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 902
diff changeset
   580
lemma term1_hom_rsp:
f7cafd3c86b0 Statement of term1_hom_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 902
diff changeset
   581
  "\<lbrakk>(alpha ===> alpha ===> op =) f_app f_app; ((op = ===> alpha) ===> op =) f_lam f_lam\<rbrakk>
f7cafd3c86b0 Statement of term1_hom_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 902
diff changeset
   582
       \<Longrightarrow> (alpha ===> op =) (term1_hom f_var f_app f_lam) (term1_hom f_var f_app f_lam)"
918
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 916
diff changeset
   583
apply(simp)
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 916
diff changeset
   584
apply(rule allI)+
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 916
diff changeset
   585
apply(rule impI)
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 916
diff changeset
   586
apply(erule alpha.induct)
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 916
diff changeset
   587
apply(auto)[1]
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 916
diff changeset
   588
apply(auto)[1]
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 916
diff changeset
   589
apply(simp)
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 916
diff changeset
   590
apply(erule conjE)+
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 916
diff changeset
   591
apply(erule exE)+
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 916
diff changeset
   592
apply(erule conjE)+
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 916
diff changeset
   593
apply(clarify)
903
f7cafd3c86b0 Statement of term1_hom_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 902
diff changeset
   594
sorry
f7cafd3c86b0 Statement of term1_hom_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 902
diff changeset
   595
f7cafd3c86b0 Statement of term1_hom_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 902
diff changeset
   596
lemma hom: "
900
3bd2847cfda7 A version of hom with quantifiers.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 898
diff changeset
   597
\<forall>f_var. \<forall>f_app \<in> Respects(alpha ===> alpha ===> op =).
3bd2847cfda7 A version of hom with quantifiers.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 898
diff changeset
   598
\<forall>f_lam \<in> Respects((op = ===> alpha) ===> op =).
3bd2847cfda7 A version of hom with quantifiers.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 898
diff changeset
   599
\<exists>hom\<in>Respects (alpha ===> op =). 
3bd2847cfda7 A version of hom with quantifiers.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 898
diff changeset
   600
    ((\<forall>x. hom (rVar x) = f_var x) \<and>
3bd2847cfda7 A version of hom with quantifiers.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 898
diff changeset
   601
     (\<forall>l r. hom (rApp l r) = f_app l r (hom l) (hom r)) \<and>
3bd2847cfda7 A version of hom with quantifiers.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 898
diff changeset
   602
     (\<forall>x a. hom (rLam a x) = f_lam (\<lambda>b. ([(a,b)]\<bullet> x)) (\<lambda>b. hom ([(a,b)] \<bullet> x))))"
3bd2847cfda7 A version of hom with quantifiers.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 898
diff changeset
   603
apply(rule allI)
3bd2847cfda7 A version of hom with quantifiers.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 898
diff changeset
   604
apply(rule ballI)+
901
28e084a66c7f term1_hom as a function
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 900
diff changeset
   605
apply(rule_tac x="term1_hom f_var f_app f_lam" in bexI)
903
f7cafd3c86b0 Statement of term1_hom_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 902
diff changeset
   606
apply(simp_all)
f7cafd3c86b0 Statement of term1_hom_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 902
diff changeset
   607
apply(simp only: in_respects)
f7cafd3c86b0 Statement of term1_hom_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 902
diff changeset
   608
apply(rule term1_hom_rsp)
f7cafd3c86b0 Statement of term1_hom_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 902
diff changeset
   609
apply(assumption)+
f7cafd3c86b0 Statement of term1_hom_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 902
diff changeset
   610
done
891
7bac7dffadeb hom lifted to hom', so it is true. Infrastructure for partially regularized quantifiers. Nicer errors for regularize.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 890
diff changeset
   611
889
cff21786d952 Appropriate respects and a statement of the lifted hom lemma
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 888
diff changeset
   612
lemma hom':
903
f7cafd3c86b0 Statement of term1_hom_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 902
diff changeset
   613
"\<exists>hom.
894
1d80641a4302 tried to witness the hom-lemma with the recursion combinator from rlam....does not work yet completely
Christian Urban <urbanc@in.tum.de>
parents: 891
diff changeset
   614
  ((\<forall>x. hom (Var x) = f_var x) \<and>
1d80641a4302 tried to witness the hom-lemma with the recursion combinator from rlam....does not work yet completely
Christian Urban <urbanc@in.tum.de>
parents: 891
diff changeset
   615
   (\<forall>l r. hom (App l r) = f_app l r (hom l) (hom r)) \<and>
1d80641a4302 tried to witness the hom-lemma with the recursion combinator from rlam....does not work yet completely
Christian Urban <urbanc@in.tum.de>
parents: 891
diff changeset
   616
   (\<forall>x a. hom (Lam a x) = f_lam (\<lambda>b. ([(a,b)] \<bullet> x)) (\<lambda>b. hom ([(a,b)] \<bullet> x))))"
903
f7cafd3c86b0 Statement of term1_hom_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 902
diff changeset
   617
apply (lifting hom)
891
7bac7dffadeb hom lifted to hom', so it is true. Infrastructure for partially regularized quantifiers. Nicer errors for regularize.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 890
diff changeset
   618
done
890
0f920b62fb7b slight tuning of relation_error
Christian Urban <urbanc@in.tum.de>
parents: 889
diff changeset
   619
896
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   620
(* test test
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   621
lemma raw_hom_correct: 
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   622
  assumes f1: "f_var \<in> Respects (op= ===> op=)"
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   623
  and     f2: "f_app \<in> Respects (alpha ===> alpha ===> op= ===> op= ===> op=)"
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   624
  and     f3: "f_lam \<in> Respects ((op= ===> alpha) ===> (op= ===> op=) ===> op=)"
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   625
  shows "\<exists>!hom\<in>Respects (alpha ===> op =). 
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   626
    ((\<forall>x. hom (rVar x) = f_var x) \<and>
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   627
     (\<forall>l r. hom (rApp l r) = f_app l r (hom l) (hom r)) \<and>
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   628
     (\<forall>x a. hom (rLam a x) = f_lam (\<lambda>b. ([(a,b)]\<bullet> x)) (\<lambda>b. hom ([(a,b)] \<bullet> x))))"
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   629
unfolding Bex1_def
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   630
apply(rule ex1I)
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   631
sorry
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 895
diff changeset
   632
*)
891
7bac7dffadeb hom lifted to hom', so it is true. Infrastructure for partially regularized quantifiers. Nicer errors for regularize.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 890
diff changeset
   633
887
d2660637e764 Incorrect version of the homomorphism lemma
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 884
diff changeset
   634
663
0dd10a900cae Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 636
diff changeset
   635
end
891
7bac7dffadeb hom lifted to hom', so it is true. Infrastructure for partially regularized quantifiers. Nicer errors for regularize.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 890
diff changeset
   636