Nominal/Term5.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Fri, 26 Feb 2010 15:10:55 +0100
changeset 1278 8814494fe4da
parent 1277 6eacf60ce41d
child 1286 87894b5156f4
permissions -rw-r--r--
Change in signature of prove_const_rsp for general lifting.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1270
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     1
theory Term5
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     2
imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" "Abs" "Perm" "Fv" "Rsp" "../Attic/Prove"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     3
begin
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     4
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     5
atom_decl name
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     6
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     7
datatype rtrm5 =
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     8
  rVr5 "name"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     9
| rAp5 "rtrm5" "rtrm5"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    10
| rLt5 "rlts" "rtrm5" --"bind (bv5 lts) in (rtrm5)"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    11
and rlts =
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    12
  rLnil
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    13
| rLcons "name" "rtrm5" "rlts"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    14
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    15
primrec
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    16
  rbv5
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    17
where
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    18
  "rbv5 rLnil = {}"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    19
| "rbv5 (rLcons n t ltl) = {atom n} \<union> (rbv5 ltl)"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    20
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    21
1277
6eacf60ce41d Permutation and FV_Alpha interface change.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1270
diff changeset
    22
setup {* snd o define_raw_perms (Datatype.the_info @{theory} "Term5.rtrm5") 2 *}
1270
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    23
print_theorems
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    24
1277
6eacf60ce41d Permutation and FV_Alpha interface change.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1270
diff changeset
    25
local_setup {* snd o define_fv_alpha (Datatype.the_info @{theory} "Term5.rtrm5") [
1270
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    26
  [[[]], [[], []], [[(SOME @{term rbv5}, 0)], [(SOME @{term rbv5}, 0)]]], [[], [[], [], []]]  ] *}
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    27
print_theorems
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    28
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    29
(* Alternate version with additional binding of name in rlts in rLcons *)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    30
(*local_setup {* snd o define_fv_alpha "Term5.rtrm5" [
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    31
  [[[]], [[], []], [[(SOME @{term rbv5}, 0)], [(SOME @{term rbv5}, 0)]]], [[], [[(NONE,0)], [], [(NONE,0)]]]  ] *}
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    32
print_theorems*)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    33
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    34
notation
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    35
  alpha_rtrm5 ("_ \<approx>5 _" [100, 100] 100) and
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    36
  alpha_rlts ("_ \<approx>l _" [100, 100] 100)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    37
thm alpha_rtrm5_alpha_rlts.intros
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    38
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    39
local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha5_inj}, []), (build_alpha_inj @{thms alpha_rtrm5_alpha_rlts.intros} @{thms rtrm5.distinct rtrm5.inject rlts.distinct rlts.inject} @{thms alpha_rtrm5.cases alpha_rlts.cases} ctxt)) ctxt)) *}
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    40
thm alpha5_inj
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    41
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    42
lemma rbv5_eqvt[eqvt]:
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    43
  "pi \<bullet> (rbv5 x) = rbv5 (pi \<bullet> x)"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    44
  apply (induct x)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    45
  apply (simp_all add: eqvts atom_eqvt)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    46
  done
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    47
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    48
lemma fv_rtrm5_rlts_eqvt[eqvt]:
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    49
  "pi \<bullet> (fv_rtrm5 x) = fv_rtrm5 (pi \<bullet> x)"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    50
  "pi \<bullet> (fv_rlts l) = fv_rlts (pi \<bullet> l)"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    51
  apply (induct x and l)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    52
  apply (simp_all add: eqvts atom_eqvt)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    53
  done
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    54
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    55
lemma alpha5_eqvt:
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    56
  "xa \<approx>5 y \<Longrightarrow> (x \<bullet> xa) \<approx>5 (x \<bullet> y)"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    57
  "xb \<approx>l ya \<Longrightarrow> (x \<bullet> xb) \<approx>l (x \<bullet> ya)"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    58
  apply (induct rule: alpha_rtrm5_alpha_rlts.inducts)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    59
  apply (simp_all add: alpha5_inj)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    60
  apply (tactic {* 
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    61
    ALLGOALS (
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    62
      TRY o REPEAT_ALL_NEW (CHANGED o rtac conjI) THEN_ALL_NEW
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    63
      (etac @{thm alpha_gen_compose_eqvt})
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    64
    ) *})
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    65
  apply (simp_all only: eqvts atom_eqvt)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    66
  done
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    67
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    68
local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha5_equivp}, []),
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    69
  (build_equivps [@{term alpha_rtrm5}, @{term alpha_rlts}] @{thm rtrm5_rlts.induct} @{thm alpha_rtrm5_alpha_rlts.induct} @{thms rtrm5.inject rlts.inject} @{thms alpha5_inj} @{thms rtrm5.distinct rlts.distinct} @{thms alpha_rtrm5.cases alpha_rlts.cases} @{thms alpha5_eqvt} ctxt)) ctxt)) *}
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    70
thm alpha5_equivp
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    71
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    72
quotient_type
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    73
  trm5 = rtrm5 / alpha_rtrm5
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    74
and
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    75
  lts = rlts / alpha_rlts
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    76
  by (auto intro: alpha5_equivp)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    77
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    78
local_setup {*
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    79
(fn ctxt => ctxt
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    80
 |> snd o (Quotient_Def.quotient_lift_const ("Vr5", @{term rVr5}))
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    81
 |> snd o (Quotient_Def.quotient_lift_const ("Ap5", @{term rAp5}))
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    82
 |> snd o (Quotient_Def.quotient_lift_const ("Lt5", @{term rLt5}))
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    83
 |> snd o (Quotient_Def.quotient_lift_const ("Lnil", @{term rLnil}))
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    84
 |> snd o (Quotient_Def.quotient_lift_const ("Lcons", @{term rLcons}))
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    85
 |> snd o (Quotient_Def.quotient_lift_const ("fv_trm5", @{term fv_rtrm5}))
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    86
 |> snd o (Quotient_Def.quotient_lift_const ("fv_lts", @{term fv_rlts}))
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    87
 |> snd o (Quotient_Def.quotient_lift_const ("bv5", @{term rbv5})))
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    88
*}
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    89
print_theorems
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    90
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    91
lemma alpha5_rfv:
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    92
  "(t \<approx>5 s \<Longrightarrow> fv_rtrm5 t = fv_rtrm5 s)"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    93
  "(l \<approx>l m \<Longrightarrow> fv_rlts l = fv_rlts m)"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    94
  apply(induct rule: alpha_rtrm5_alpha_rlts.inducts)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    95
  apply(simp_all add: alpha_gen)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    96
  done
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    97
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    98
lemma bv_list_rsp:
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    99
  shows "x \<approx>l y \<Longrightarrow> rbv5 x = rbv5 y"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   100
  apply(induct rule: alpha_rtrm5_alpha_rlts.inducts(2))
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   101
  apply(simp_all)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   102
  done
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   103
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   104
lemma [quot_respect]:
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   105
  "(alpha_rlts ===> op =) fv_rlts fv_rlts"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   106
  "(alpha_rtrm5 ===> op =) fv_rtrm5 fv_rtrm5"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   107
  "(alpha_rlts ===> op =) rbv5 rbv5"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   108
  "(op = ===> alpha_rtrm5) rVr5 rVr5"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   109
  "(alpha_rtrm5 ===> alpha_rtrm5 ===> alpha_rtrm5) rAp5 rAp5"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   110
  "(alpha_rlts ===> alpha_rtrm5 ===> alpha_rtrm5) rLt5 rLt5"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   111
  "(op = ===> alpha_rtrm5 ===> alpha_rlts ===> alpha_rlts) rLcons rLcons"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   112
  "(op = ===> alpha_rtrm5 ===> alpha_rtrm5) permute permute"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   113
  "(op = ===> alpha_rlts ===> alpha_rlts) permute permute"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   114
  apply (simp_all add: alpha5_inj alpha5_rfv alpha5_eqvt bv_list_rsp)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   115
  apply (clarify) apply (rule conjI)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   116
  apply (rule_tac x="0" in exI) apply (simp add: fresh_star_def fresh_zero_perm alpha_gen alpha5_rfv)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   117
  apply (rule_tac x="0" in exI) apply (simp add: fresh_star_def fresh_zero_perm alpha_gen alpha5_rfv)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   118
  done
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   119
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   120
lemma
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   121
  shows "(alpha_rlts ===> op =) rbv5 rbv5"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   122
  by (simp add: bv_list_rsp)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   123
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   124
lemmas trm5_lts_inducts = rtrm5_rlts.inducts[quot_lifted]
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   125
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   126
instantiation trm5 and lts :: pt
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   127
begin
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   128
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   129
quotient_definition
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   130
  "permute_trm5 :: perm \<Rightarrow> trm5 \<Rightarrow> trm5"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   131
is
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   132
  "permute :: perm \<Rightarrow> rtrm5 \<Rightarrow> rtrm5"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   133
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   134
quotient_definition
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   135
  "permute_lts :: perm \<Rightarrow> lts \<Rightarrow> lts"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   136
is
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   137
  "permute :: perm \<Rightarrow> rlts \<Rightarrow> rlts"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   138
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   139
instance by default
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   140
  (simp_all add: permute_rtrm5_permute_rlts_zero[quot_lifted] permute_rtrm5_permute_rlts_append[quot_lifted])
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   141
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   142
end
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   143
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   144
lemmas
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   145
    permute_trm5_lts = permute_rtrm5_permute_rlts.simps[quot_lifted]
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   146
and alpha5_INJ = alpha5_inj[unfolded alpha_gen, quot_lifted, folded alpha_gen]
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   147
and bv5[simp] = rbv5.simps[quot_lifted]
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   148
and fv_trm5_lts[simp] = fv_rtrm5_fv_rlts.simps[quot_lifted]
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   149
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   150
lemma lets_ok:
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   151
  "(Lt5 (Lcons x (Vr5 x) Lnil) (Vr5 x)) = (Lt5 (Lcons y (Vr5 y) Lnil) (Vr5 y))"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   152
apply (subst alpha5_INJ)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   153
apply (rule conjI)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   154
apply (rule_tac x="(x \<leftrightarrow> y)" in exI)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   155
apply (simp only: alpha_gen)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   156
apply (simp add: permute_trm5_lts fresh_star_def)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   157
apply (rule_tac x="(x \<leftrightarrow> y)" in exI)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   158
apply (simp only: alpha_gen)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   159
apply (simp add: permute_trm5_lts fresh_star_def)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   160
done
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   161
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   162
lemma lets_ok2:
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   163
  "(Lt5 (Lcons x (Vr5 x) (Lcons y (Vr5 y) Lnil)) (Ap5 (Vr5 x) (Vr5 y))) =
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   164
   (Lt5 (Lcons y (Vr5 y) (Lcons x (Vr5 x) Lnil)) (Ap5 (Vr5 x) (Vr5 y)))"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   165
apply (subst alpha5_INJ)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   166
apply (rule conjI)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   167
apply (rule_tac x="(x \<leftrightarrow> y)" in exI)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   168
apply (simp only: alpha_gen)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   169
apply (simp add: permute_trm5_lts fresh_star_def)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   170
apply (rule_tac x="0 :: perm" in exI)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   171
apply (simp only: alpha_gen)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   172
apply (simp add: permute_trm5_lts fresh_star_def)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   173
done
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   174
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   175
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   176
lemma lets_not_ok1:
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   177
  "x \<noteq> y \<Longrightarrow> (Lt5 (Lcons x (Vr5 x) (Lcons y (Vr5 y) Lnil)) (Ap5 (Vr5 x) (Vr5 y))) \<noteq>
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   178
             (Lt5 (Lcons y (Vr5 x) (Lcons x (Vr5 y) Lnil)) (Ap5 (Vr5 x) (Vr5 y)))"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   179
apply (simp add: alpha5_INJ(3) alpha_gen)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   180
apply (simp add: permute_trm5_lts fresh_star_def alpha5_INJ(5) alpha5_INJ(2) alpha5_INJ(1))
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   181
done
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   182
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   183
lemma distinct_helper:
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   184
  shows "\<not>(rVr5 x \<approx>5 rAp5 y z)"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   185
  apply auto
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   186
  apply (erule alpha_rtrm5.cases)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   187
  apply (simp_all only: rtrm5.distinct)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   188
  done
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   189
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   190
lemma distinct_helper2:
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   191
  shows "(Vr5 x) \<noteq> (Ap5 y z)"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   192
  by (lifting distinct_helper)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   193
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   194
lemma lets_nok:
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   195
  "x \<noteq> y \<Longrightarrow> x \<noteq> z \<Longrightarrow> z \<noteq> y \<Longrightarrow>
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   196
   (Lt5 (Lcons x (Ap5 (Vr5 z) (Vr5 z)) (Lcons y (Vr5 z) Lnil)) (Ap5 (Vr5 x) (Vr5 y))) \<noteq>
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   197
   (Lt5 (Lcons y (Vr5 z) (Lcons x (Ap5 (Vr5 z) (Vr5 z)) Lnil)) (Ap5 (Vr5 x) (Vr5 y)))"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   198
apply (simp only: alpha5_INJ(3) alpha5_INJ(5) alpha_gen permute_trm5_lts fresh_star_def)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   199
apply (simp add: distinct_helper2)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   200
done
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   201
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   202
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   203
end