src/HOL/Types_To_Sets/Examples/Prerequisites.thy
author wenzelm
Sat, 05 Jan 2019 17:24:33 +0100
changeset 69597 ff784d5a5bfb
parent 69296 bc0b0d465991
child 69912 dd55d2c926d9
permissions -rw-r--r--
isabelle update -u control_cartouches;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
64551
79e9587dbcca proper session HOL-Types_To_Sets;
wenzelm
parents:
diff changeset
     1
(*  Title:      HOL/Types_To_Sets/Examples/Prerequisites.thy
79e9587dbcca proper session HOL-Types_To_Sets;
wenzelm
parents:
diff changeset
     2
    Author:     Ondřej Kunčar, TU München
79e9587dbcca proper session HOL-Types_To_Sets;
wenzelm
parents:
diff changeset
     3
*)
79e9587dbcca proper session HOL-Types_To_Sets;
wenzelm
parents:
diff changeset
     4
79e9587dbcca proper session HOL-Types_To_Sets;
wenzelm
parents:
diff changeset
     5
theory Prerequisites
79e9587dbcca proper session HOL-Types_To_Sets;
wenzelm
parents:
diff changeset
     6
  imports Main
69296
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
     7
  keywords "lemmas_with"::thy_decl
64551
79e9587dbcca proper session HOL-Types_To_Sets;
wenzelm
parents:
diff changeset
     8
begin
79e9587dbcca proper session HOL-Types_To_Sets;
wenzelm
parents:
diff changeset
     9
79e9587dbcca proper session HOL-Types_To_Sets;
wenzelm
parents:
diff changeset
    10
context
79e9587dbcca proper session HOL-Types_To_Sets;
wenzelm
parents:
diff changeset
    11
  fixes Rep Abs A T
79e9587dbcca proper session HOL-Types_To_Sets;
wenzelm
parents:
diff changeset
    12
  assumes type: "type_definition Rep Abs A"
79e9587dbcca proper session HOL-Types_To_Sets;
wenzelm
parents:
diff changeset
    13
  assumes T_def: "T \<equiv> (\<lambda>(x::'a) (y::'b). x = Rep y)"
79e9587dbcca proper session HOL-Types_To_Sets;
wenzelm
parents:
diff changeset
    14
begin
79e9587dbcca proper session HOL-Types_To_Sets;
wenzelm
parents:
diff changeset
    15
79e9587dbcca proper session HOL-Types_To_Sets;
wenzelm
parents:
diff changeset
    16
lemma type_definition_Domainp: "Domainp T = (\<lambda>x. x \<in> A)"
79e9587dbcca proper session HOL-Types_To_Sets;
wenzelm
parents:
diff changeset
    17
proof -
79e9587dbcca proper session HOL-Types_To_Sets;
wenzelm
parents:
diff changeset
    18
  interpret type_definition Rep Abs A by(rule type)
79e9587dbcca proper session HOL-Types_To_Sets;
wenzelm
parents:
diff changeset
    19
  show ?thesis
79e9587dbcca proper session HOL-Types_To_Sets;
wenzelm
parents:
diff changeset
    20
    unfolding Domainp_iff[abs_def] T_def fun_eq_iff
79e9587dbcca proper session HOL-Types_To_Sets;
wenzelm
parents:
diff changeset
    21
    by (metis Abs_inverse Rep)
79e9587dbcca proper session HOL-Types_To_Sets;
wenzelm
parents:
diff changeset
    22
qed
79e9587dbcca proper session HOL-Types_To_Sets;
wenzelm
parents:
diff changeset
    23
79e9587dbcca proper session HOL-Types_To_Sets;
wenzelm
parents:
diff changeset
    24
end
79e9587dbcca proper session HOL-Types_To_Sets;
wenzelm
parents:
diff changeset
    25
69295
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    26
subsection \<open>setting up transfer for local typedef\<close>
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    27
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    28
lemmas [transfer_rule] = \<comment> \<open>prefer right-total rules\<close>
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    29
  right_total_All_transfer
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    30
  right_total_UNIV_transfer
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    31
  right_total_Ex_transfer
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    32
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    33
locale local_typedef = fixes S ::"'b set" and s::"'s itself"
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    34
  assumes Ex_type_definition_S: "\<exists>(Rep::'s \<Rightarrow> 'b) (Abs::'b \<Rightarrow> 's). type_definition Rep Abs S"
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    35
begin
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    36
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    37
definition "rep = fst (SOME (Rep::'s \<Rightarrow> 'b, Abs). type_definition Rep Abs S)"
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    38
definition "Abs = snd (SOME (Rep::'s \<Rightarrow> 'b, Abs). type_definition Rep Abs S)"
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    39
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    40
lemma type_definition_S: "type_definition rep Abs S"
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    41
  unfolding Abs_def rep_def split_beta'
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    42
  by (rule someI_ex) (use Ex_type_definition_S in auto)
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    43
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    44
lemma rep_in_S[simp]: "rep x \<in> S"
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    45
  and rep_inverse[simp]: "Abs (rep x) = x"
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    46
  and Abs_inverse[simp]: "y \<in> S \<Longrightarrow> rep (Abs y) = y"
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    47
  using type_definition_S
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    48
  unfolding type_definition_def by auto
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    49
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    50
definition cr_S where "cr_S \<equiv> \<lambda>s b. s = rep b"
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    51
lemmas Domainp_cr_S = type_definition_Domainp[OF type_definition_S cr_S_def, transfer_domain_rule]
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    52
lemmas right_total_cr_S = typedef_right_total[OF type_definition_S cr_S_def, transfer_rule]
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    53
  and bi_unique_cr_S = typedef_bi_unique[OF type_definition_S cr_S_def, transfer_rule]
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    54
  and left_unique_cr_S = typedef_left_unique[OF type_definition_S cr_S_def, transfer_rule]
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    55
  and right_unique_cr_S = typedef_right_unique[OF type_definition_S cr_S_def, transfer_rule]
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    56
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    57
lemma cr_S_rep[intro, simp]: "cr_S (rep a) a" by (simp add: cr_S_def)
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    58
lemma cr_S_Abs[intro, simp]: "a\<in>S \<Longrightarrow> cr_S a (Abs a)" by (simp add: cr_S_def)
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    59
64551
79e9587dbcca proper session HOL-Types_To_Sets;
wenzelm
parents:
diff changeset
    60
end
69295
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    61
69296
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    62
subsection \<open>some \<close>
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    63
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    64
subsection \<open>Tool support\<close>
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    65
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    66
lemmas subset_iff' = subset_iff[folded Ball_def]
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    67
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    68
ML \<open>
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    69
structure More_Simplifier =
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    70
struct
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    71
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    72
fun asm_full_var_simplify ctxt thm =
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    73
  let
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    74
    val ((_, [thm']), ctxt') = Variable.import false [thm] ctxt
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    75
  in
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    76
    Simplifier.asm_full_simplify ctxt' thm'
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    77
    |> singleton (Variable.export ctxt' ctxt)
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    78
    |> Drule.zero_var_indexes
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    79
  end
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    80
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    81
fun var_simplify_only ctxt ths thm =
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    82
  asm_full_var_simplify (Raw_Simplifier.clear_simpset ctxt addsimps ths) thm
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    83
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    84
val var_simplified = Attrib.thms >>
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    85
  (fn ths => Thm.rule_attribute ths
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    86
    (fn context => var_simplify_only (Context.proof_of context) ths))
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    87
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    88
val _ = Theory.setup (Attrib.setup \<^binding>\<open>var_simplified\<close> var_simplified "simplified rule (with vars)")
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    89
69295
b8b33ef47f40 use locales in Group_On_With
immler
parents: 64551
diff changeset
    90
end
69296
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    91
\<close>
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    92
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    93
ML \<open>
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    94
val _ = Outer_Syntax.local_theory' \<^command_keyword>\<open>lemmas_with\<close> "note theorems with (the same) attributes"
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69296
diff changeset
    95
    (Parse.attribs --| \<^keyword>\<open>:\<close> -- Parse_Spec.name_facts -- Parse.for_fixes
69296
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    96
     >> (fn (((attrs),facts), fixes) =>
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    97
      #2 oo Specification.theorems_cmd Thm.theoremK
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    98
        (map (apsnd (map (apsnd (fn xs => attrs@xs)))) facts) fixes))
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
    99
\<close>
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
   100
bc0b0d465991 generalized local_typedef_ab_group_add
immler
parents: 69295
diff changeset
   101
end