src/HOL/HOLCF/One.thy
author paulson <lp15@cam.ac.uk>
Sun, 14 Aug 2022 23:51:47 +0100
changeset 75864 3842556b757c
parent 69597 ff784d5a5bfb
permissions -rw-r--r--
moved some material from Sum_of_Powers
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42151
4da4fc77664b tuned headers;
wenzelm
parents: 41295
diff changeset
     1
(*  Title:      HOL/HOLCF/One.thy
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2275
diff changeset
     2
    Author:     Oscar Slotosch
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     3
*)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     4
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
     5
section \<open>The unit domain\<close>
15577
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
     6
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
     7
theory One
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
     8
  imports Lift
15577
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
     9
begin
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    10
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    11
type_synonym one = "unit lift"
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    12
41295
5b5388d4ccc9 types -> type_synonym
huffman
parents: 41182
diff changeset
    13
translations
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    14
  (type) "one" \<leftharpoondown> (type) "unit lift"
41295
5b5388d4ccc9 types -> type_synonym
huffman
parents: 41182
diff changeset
    15
5b5388d4ccc9 types -> type_synonym
huffman
parents: 41182
diff changeset
    16
definition ONE :: "one"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    17
  where "ONE \<equiv> Def ()"
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    18
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67312
diff changeset
    19
text \<open>Exhaustion and Elimination for type \<^typ>\<open>one\<close>\<close>
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    20
16747
934b6b36d794 cleaned up
huffman
parents: 16070
diff changeset
    21
lemma Exh_one: "t = \<bottom> \<or> t = ONE"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    22
  by (induct t) (simp_all add: ONE_def)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    23
35783
38538bfe9ca6 declare case_names for various induction rules
huffman
parents: 35431
diff changeset
    24
lemma oneE [case_names bottom ONE]: "\<lbrakk>p = \<bottom> \<Longrightarrow> Q; p = ONE \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    25
  by (induct p) (simp_all add: ONE_def)
27293
de9a2fd0eab4 added some lemmas; tuned proofs
huffman
parents: 25131
diff changeset
    26
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    27
lemma one_induct [case_names bottom ONE]: "P \<bottom> \<Longrightarrow> P ONE \<Longrightarrow> P x"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    28
  by (cases x rule: oneE) simp_all
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    29
41182
717404c7d59a add notsqsubseteq syntax
huffman
parents: 40774
diff changeset
    30
lemma dist_below_one [simp]: "ONE \<notsqsubseteq> \<bottom>"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    31
  by (simp add: ONE_def)
27293
de9a2fd0eab4 added some lemmas; tuned proofs
huffman
parents: 25131
diff changeset
    32
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 30911
diff changeset
    33
lemma below_ONE [simp]: "x \<sqsubseteq> ONE"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    34
  by (induct x rule: one_induct) simp_all
27293
de9a2fd0eab4 added some lemmas; tuned proofs
huffman
parents: 25131
diff changeset
    35
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 30911
diff changeset
    36
lemma ONE_below_iff [simp]: "ONE \<sqsubseteq> x \<longleftrightarrow> x = ONE"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    37
  by (induct x rule: one_induct) simp_all
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    38
30911
7809cbaa1b61 domain package: simplify internal proofs of con_rews
huffman
parents: 29141
diff changeset
    39
lemma ONE_defined [simp]: "ONE \<noteq> \<bottom>"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    40
  by (simp add: ONE_def)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    41
27293
de9a2fd0eab4 added some lemmas; tuned proofs
huffman
parents: 25131
diff changeset
    42
lemma one_neq_iffs [simp]:
de9a2fd0eab4 added some lemmas; tuned proofs
huffman
parents: 25131
diff changeset
    43
  "x \<noteq> ONE \<longleftrightarrow> x = \<bottom>"
de9a2fd0eab4 added some lemmas; tuned proofs
huffman
parents: 25131
diff changeset
    44
  "ONE \<noteq> x \<longleftrightarrow> x = \<bottom>"
de9a2fd0eab4 added some lemmas; tuned proofs
huffman
parents: 25131
diff changeset
    45
  "x \<noteq> \<bottom> \<longleftrightarrow> x = ONE"
de9a2fd0eab4 added some lemmas; tuned proofs
huffman
parents: 25131
diff changeset
    46
  "\<bottom> \<noteq> x \<longleftrightarrow> x = ONE"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    47
  by (induct x rule: one_induct) simp_all
27293
de9a2fd0eab4 added some lemmas; tuned proofs
huffman
parents: 25131
diff changeset
    48
de9a2fd0eab4 added some lemmas; tuned proofs
huffman
parents: 25131
diff changeset
    49
lemma compact_ONE: "compact ONE"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    50
  by (rule compact_chfin)
17838
3032e90c4975 added compactness theorems
huffman
parents: 16747
diff changeset
    51
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67312
diff changeset
    52
text \<open>Case analysis function for type \<^typ>\<open>one\<close>\<close>
18080
c1a7490ee3ff add constant one_when; LAM pattern for ONE
huffman
parents: 17838
diff changeset
    53
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    54
definition one_case :: "'a::pcpo \<rightarrow> one \<rightarrow> 'a"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    55
  where "one_case = (\<Lambda> a x. seq\<cdot>x\<cdot>a)"
18080
c1a7490ee3ff add constant one_when; LAM pattern for ONE
huffman
parents: 17838
diff changeset
    56
c1a7490ee3ff add constant one_when; LAM pattern for ONE
huffman
parents: 17838
diff changeset
    57
translations
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    58
  "case x of XCONST ONE \<Rightarrow> t" \<rightleftharpoons> "CONST one_case\<cdot>t\<cdot>x"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    59
  "case x of XCONST ONE :: 'a \<Rightarrow> t" \<rightharpoonup> "CONST one_case\<cdot>t\<cdot>x"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    60
  "\<Lambda> (XCONST ONE). t" \<rightleftharpoons> "CONST one_case\<cdot>t"
18080
c1a7490ee3ff add constant one_when; LAM pattern for ONE
huffman
parents: 17838
diff changeset
    61
40212
20df78048db5 rename constant 'one_when' to 'one_case'
huffman
parents: 40046
diff changeset
    62
lemma one_case1 [simp]: "(case \<bottom> of ONE \<Rightarrow> t) = \<bottom>"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    63
  by (simp add: one_case_def)
18080
c1a7490ee3ff add constant one_when; LAM pattern for ONE
huffman
parents: 17838
diff changeset
    64
40212
20df78048db5 rename constant 'one_when' to 'one_case'
huffman
parents: 40046
diff changeset
    65
lemma one_case2 [simp]: "(case ONE of ONE \<Rightarrow> t) = t"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    66
  by (simp add: one_case_def)
18080
c1a7490ee3ff add constant one_when; LAM pattern for ONE
huffman
parents: 17838
diff changeset
    67
40212
20df78048db5 rename constant 'one_when' to 'one_case'
huffman
parents: 40046
diff changeset
    68
lemma one_case3 [simp]: "(case x of ONE \<Rightarrow> ONE) = x"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    69
  by (induct x rule: one_induct) simp_all
18080
c1a7490ee3ff add constant one_when; LAM pattern for ONE
huffman
parents: 17838
diff changeset
    70
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    71
end